# HG changeset patch # User wenzelm # Date 1332752216 -7200 # Node ID f067afe98049e4ca183ecb00774a16e2ef52505b # Parent 29e92b644d6c4ff4c910e5cf3dca3b99c0861c3f# Parent db502663179986e8be3a29d8dbbc983747a130e2 merged, resolving trivial conflict; diff -r 29e92b644d6c -r f067afe98049 Admin/isatest/isatest-settings --- a/Admin/isatest/isatest-settings Fri Mar 23 20:32:43 2012 +0100 +++ b/Admin/isatest/isatest-settings Mon Mar 26 10:56:56 2012 +0200 @@ -22,7 +22,8 @@ bulwahn@in.tum.de \ hoelzl@in.tum.de \ krauss@in.tum.de \ -noschinl@in.tum.de" +noschinl@in.tum.de \ +kuncar@in.tum.de" LOGPREFIX=$HOME/log MASTERLOG=$LOGPREFIX/isatest.log diff -r 29e92b644d6c -r f067afe98049 Admin/isatest/pmail --- a/Admin/isatest/pmail Fri Mar 23 20:32:43 2012 +0100 +++ b/Admin/isatest/pmail Mon Mar 26 10:56:56 2012 +0200 @@ -95,7 +95,7 @@ case `uname` in Linux) for F in $@; do ATTACH="$ATTACH -a $F"; done - cat "$BODY" | mail -s "$SUBJECT" $ATTACH "$TO" + cat "$BODY" | mail -Ssmtp=mailbroy.informatik.tu-muenchen.de -s "$SUBJECT" $ATTACH "$TO" ;; SunOS) print_body "$SUBJECT" "$BODY" $@ | mail -t "$TO" diff -r 29e92b644d6c -r f067afe98049 NEWS --- a/NEWS Fri Mar 23 20:32:43 2012 +0100 +++ b/NEWS Mon Mar 26 10:56:56 2012 +0200 @@ -90,6 +90,30 @@ *** HOL *** +* The representation of numerals has changed. We now have a datatype +"num" representing strictly positive binary numerals, along with +functions "numeral :: num => 'a" and "neg_numeral :: num => 'a" to +represent positive and negated numeric literals, respectively. (See +definitions in Num.thy.) Potential INCOMPATIBILITY; some user theories +may require adaptations: + + - Theorems with number_ring or number_semiring constraints: These + classes are gone; use comm_ring_1 or comm_semiring_1 instead. + + - Theories defining numeric types: Remove number, number_semiring, + and number_ring instances. Defer all theorems about numerals until + after classes one and semigroup_add have been instantiated. + + - Numeral-only simp rules: Replace each rule having a "number_of v" + pattern with two copies, one for numeral and one for neg_numeral. + + - Theorems about subclasses of semiring_1 or ring_1: These classes + automatically support numerals now, so more simp rules and + simprocs may now apply within the proof. + + - Definitions and theorems using old constructors Pls/Min/Bit0/Bit1: + Redefine using other integer operations. + * Type 'a set is now a proper type constructor (just as before Isabelle2008). Definitions mem_def and Collect_def have disappeared. Non-trivial INCOMPATIBILITY. For developments keeping predicates and diff -r 29e92b644d6c -r f067afe98049 etc/isar-keywords.el --- a/etc/isar-keywords.el Fri Mar 23 20:32:43 2012 +0100 +++ b/etc/isar-keywords.el Mon Mar 26 10:56:56 2012 +0200 @@ -221,6 +221,7 @@ "sect" "section" "setup" + "setup_lifting" "show" "simproc_setup" "sledgehammer" @@ -518,13 +519,13 @@ "print_translation" "quickcheck_generator" "quickcheck_params" - "quotient_definition" "realizability" "realizers" "recdef" "record" "refute_params" "setup" + "setup_lifting" "simproc_setup" "sledgehammer_params" "spark_end" @@ -563,6 +564,7 @@ "nominal_inductive2" "nominal_primrec" "pcpodef" + "quotient_definition" "quotient_type" "recdef_tc" "rep_datatype" diff -r 29e92b644d6c -r f067afe98049 src/HOL/Algebra/Group.thy --- a/src/HOL/Algebra/Group.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Algebra/Group.thy Mon Mar 26 10:56:56 2012 +0200 @@ -30,7 +30,7 @@ where "Units G = {y. y \ carrier G & (\x \ carrier G. x \\<^bsub>G\<^esub> y = \\<^bsub>G\<^esub> & y \\<^bsub>G\<^esub> x = \\<^bsub>G\<^esub>)}" consts - pow :: "[('a, 'm) monoid_scheme, 'a, 'b::number] => 'a" (infixr "'(^')\" 75) + pow :: "[('a, 'm) monoid_scheme, 'a, 'b::semiring_1] => 'a" (infixr "'(^')\" 75) overloading nat_pow == "pow :: [_, 'a, nat] => 'a" begin diff -r 29e92b644d6c -r f067afe98049 src/HOL/Archimedean_Field.thy --- a/src/HOL/Archimedean_Field.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Archimedean_Field.thy Mon Mar 26 10:56:56 2012 +0200 @@ -12,7 +12,7 @@ text {* Archimedean fields have no infinite elements. *} -class archimedean_field = linordered_field + number_ring + +class archimedean_field = linordered_field + assumes ex_le_of_int: "\z. x \ of_int z" lemma ex_less_of_int: @@ -202,8 +202,11 @@ lemma floor_one [simp]: "floor 1 = 1" using floor_of_int [of 1] by simp -lemma floor_number_of [simp]: "floor (number_of v) = number_of v" - using floor_of_int [of "number_of v"] by simp +lemma floor_numeral [simp]: "floor (numeral v) = numeral v" + using floor_of_int [of "numeral v"] by simp + +lemma floor_neg_numeral [simp]: "floor (neg_numeral v) = neg_numeral v" + using floor_of_int [of "neg_numeral v"] by simp lemma zero_le_floor [simp]: "0 \ floor x \ 0 \ x" by (simp add: le_floor_iff) @@ -211,7 +214,12 @@ lemma one_le_floor [simp]: "1 \ floor x \ 1 \ x" by (simp add: le_floor_iff) -lemma number_of_le_floor [simp]: "number_of v \ floor x \ number_of v \ x" +lemma numeral_le_floor [simp]: + "numeral v \ floor x \ numeral v \ x" + by (simp add: le_floor_iff) + +lemma neg_numeral_le_floor [simp]: + "neg_numeral v \ floor x \ neg_numeral v \ x" by (simp add: le_floor_iff) lemma zero_less_floor [simp]: "0 < floor x \ 1 \ x" @@ -220,8 +228,12 @@ lemma one_less_floor [simp]: "1 < floor x \ 2 \ x" by (simp add: less_floor_iff) -lemma number_of_less_floor [simp]: - "number_of v < floor x \ number_of v + 1 \ x" +lemma numeral_less_floor [simp]: + "numeral v < floor x \ numeral v + 1 \ x" + by (simp add: less_floor_iff) + +lemma neg_numeral_less_floor [simp]: + "neg_numeral v < floor x \ neg_numeral v + 1 \ x" by (simp add: less_floor_iff) lemma floor_le_zero [simp]: "floor x \ 0 \ x < 1" @@ -230,8 +242,12 @@ lemma floor_le_one [simp]: "floor x \ 1 \ x < 2" by (simp add: floor_le_iff) -lemma floor_le_number_of [simp]: - "floor x \ number_of v \ x < number_of v + 1" +lemma floor_le_numeral [simp]: + "floor x \ numeral v \ x < numeral v + 1" + by (simp add: floor_le_iff) + +lemma floor_le_neg_numeral [simp]: + "floor x \ neg_numeral v \ x < neg_numeral v + 1" by (simp add: floor_le_iff) lemma floor_less_zero [simp]: "floor x < 0 \ x < 0" @@ -240,8 +256,12 @@ lemma floor_less_one [simp]: "floor x < 1 \ x < 1" by (simp add: floor_less_iff) -lemma floor_less_number_of [simp]: - "floor x < number_of v \ x < number_of v" +lemma floor_less_numeral [simp]: + "floor x < numeral v \ x < numeral v" + by (simp add: floor_less_iff) + +lemma floor_less_neg_numeral [simp]: + "floor x < neg_numeral v \ x < neg_numeral v" by (simp add: floor_less_iff) text {* Addition and subtraction of integers *} @@ -249,9 +269,13 @@ lemma floor_add_of_int [simp]: "floor (x + of_int z) = floor x + z" using floor_correct [of x] by (simp add: floor_unique) -lemma floor_add_number_of [simp]: - "floor (x + number_of v) = floor x + number_of v" - using floor_add_of_int [of x "number_of v"] by simp +lemma floor_add_numeral [simp]: + "floor (x + numeral v) = floor x + numeral v" + using floor_add_of_int [of x "numeral v"] by simp + +lemma floor_add_neg_numeral [simp]: + "floor (x + neg_numeral v) = floor x + neg_numeral v" + using floor_add_of_int [of x "neg_numeral v"] by simp lemma floor_add_one [simp]: "floor (x + 1) = floor x + 1" using floor_add_of_int [of x 1] by simp @@ -259,9 +283,13 @@ lemma floor_diff_of_int [simp]: "floor (x - of_int z) = floor x - z" using floor_add_of_int [of x "- z"] by (simp add: algebra_simps) -lemma floor_diff_number_of [simp]: - "floor (x - number_of v) = floor x - number_of v" - using floor_diff_of_int [of x "number_of v"] by simp +lemma floor_diff_numeral [simp]: + "floor (x - numeral v) = floor x - numeral v" + using floor_diff_of_int [of x "numeral v"] by simp + +lemma floor_diff_neg_numeral [simp]: + "floor (x - neg_numeral v) = floor x - neg_numeral v" + using floor_diff_of_int [of x "neg_numeral v"] by simp lemma floor_diff_one [simp]: "floor (x - 1) = floor x - 1" using floor_diff_of_int [of x 1] by simp @@ -320,8 +348,11 @@ lemma ceiling_one [simp]: "ceiling 1 = 1" using ceiling_of_int [of 1] by simp -lemma ceiling_number_of [simp]: "ceiling (number_of v) = number_of v" - using ceiling_of_int [of "number_of v"] by simp +lemma ceiling_numeral [simp]: "ceiling (numeral v) = numeral v" + using ceiling_of_int [of "numeral v"] by simp + +lemma ceiling_neg_numeral [simp]: "ceiling (neg_numeral v) = neg_numeral v" + using ceiling_of_int [of "neg_numeral v"] by simp lemma ceiling_le_zero [simp]: "ceiling x \ 0 \ x \ 0" by (simp add: ceiling_le_iff) @@ -329,8 +360,12 @@ lemma ceiling_le_one [simp]: "ceiling x \ 1 \ x \ 1" by (simp add: ceiling_le_iff) -lemma ceiling_le_number_of [simp]: - "ceiling x \ number_of v \ x \ number_of v" +lemma ceiling_le_numeral [simp]: + "ceiling x \ numeral v \ x \ numeral v" + by (simp add: ceiling_le_iff) + +lemma ceiling_le_neg_numeral [simp]: + "ceiling x \ neg_numeral v \ x \ neg_numeral v" by (simp add: ceiling_le_iff) lemma ceiling_less_zero [simp]: "ceiling x < 0 \ x \ -1" @@ -339,8 +374,12 @@ lemma ceiling_less_one [simp]: "ceiling x < 1 \ x \ 0" by (simp add: ceiling_less_iff) -lemma ceiling_less_number_of [simp]: - "ceiling x < number_of v \ x \ number_of v - 1" +lemma ceiling_less_numeral [simp]: + "ceiling x < numeral v \ x \ numeral v - 1" + by (simp add: ceiling_less_iff) + +lemma ceiling_less_neg_numeral [simp]: + "ceiling x < neg_numeral v \ x \ neg_numeral v - 1" by (simp add: ceiling_less_iff) lemma zero_le_ceiling [simp]: "0 \ ceiling x \ -1 < x" @@ -349,8 +388,12 @@ lemma one_le_ceiling [simp]: "1 \ ceiling x \ 0 < x" by (simp add: le_ceiling_iff) -lemma number_of_le_ceiling [simp]: - "number_of v \ ceiling x\ number_of v - 1 < x" +lemma numeral_le_ceiling [simp]: + "numeral v \ ceiling x \ numeral v - 1 < x" + by (simp add: le_ceiling_iff) + +lemma neg_numeral_le_ceiling [simp]: + "neg_numeral v \ ceiling x \ neg_numeral v - 1 < x" by (simp add: le_ceiling_iff) lemma zero_less_ceiling [simp]: "0 < ceiling x \ 0 < x" @@ -359,8 +402,12 @@ lemma one_less_ceiling [simp]: "1 < ceiling x \ 1 < x" by (simp add: less_ceiling_iff) -lemma number_of_less_ceiling [simp]: - "number_of v < ceiling x \ number_of v < x" +lemma numeral_less_ceiling [simp]: + "numeral v < ceiling x \ numeral v < x" + by (simp add: less_ceiling_iff) + +lemma neg_numeral_less_ceiling [simp]: + "neg_numeral v < ceiling x \ neg_numeral v < x" by (simp add: less_ceiling_iff) text {* Addition and subtraction of integers *} @@ -368,9 +415,13 @@ lemma ceiling_add_of_int [simp]: "ceiling (x + of_int z) = ceiling x + z" using ceiling_correct [of x] by (simp add: ceiling_unique) -lemma ceiling_add_number_of [simp]: - "ceiling (x + number_of v) = ceiling x + number_of v" - using ceiling_add_of_int [of x "number_of v"] by simp +lemma ceiling_add_numeral [simp]: + "ceiling (x + numeral v) = ceiling x + numeral v" + using ceiling_add_of_int [of x "numeral v"] by simp + +lemma ceiling_add_neg_numeral [simp]: + "ceiling (x + neg_numeral v) = ceiling x + neg_numeral v" + using ceiling_add_of_int [of x "neg_numeral v"] by simp lemma ceiling_add_one [simp]: "ceiling (x + 1) = ceiling x + 1" using ceiling_add_of_int [of x 1] by simp @@ -378,9 +429,13 @@ lemma ceiling_diff_of_int [simp]: "ceiling (x - of_int z) = ceiling x - z" using ceiling_add_of_int [of x "- z"] by (simp add: algebra_simps) -lemma ceiling_diff_number_of [simp]: - "ceiling (x - number_of v) = ceiling x - number_of v" - using ceiling_diff_of_int [of x "number_of v"] by simp +lemma ceiling_diff_numeral [simp]: + "ceiling (x - numeral v) = ceiling x - numeral v" + using ceiling_diff_of_int [of x "numeral v"] by simp + +lemma ceiling_diff_neg_numeral [simp]: + "ceiling (x - neg_numeral v) = ceiling x - neg_numeral v" + using ceiling_diff_of_int [of x "neg_numeral v"] by simp lemma ceiling_diff_one [simp]: "ceiling (x - 1) = ceiling x - 1" using ceiling_diff_of_int [of x 1] by simp diff -r 29e92b644d6c -r f067afe98049 src/HOL/Code_Evaluation.thy --- a/src/HOL/Code_Evaluation.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Code_Evaluation.thy Mon Mar 26 10:56:56 2012 +0200 @@ -146,33 +146,29 @@ "term_of_num_semiring two = (\_. dummy_term)" lemma (in term_syntax) term_of_num_semiring_code [code]: - "term_of_num_semiring two k = (if k = 0 then termify Int.Pls + "term_of_num_semiring two k = ( + if k = 1 then termify Num.One else (if k mod two = 0 - 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) + then termify Num.Bit0 <\> term_of_num_semiring two (k div two) + else termify Num.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) lemma (in term_syntax) term_of_nat_code [code]: - "term_of (n::nat) = termify (number_of :: int \ nat) <\> term_of_num_semiring (2::nat) n" + "term_of (n::nat) = ( + if n = 0 then termify (0 :: nat) + else termify (numeral :: num \ 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_semiring (2::code_numeral) k" + "term_of (k::code_numeral) = ( + if k = 0 then termify (0 :: code_numeral) + else termify (numeral :: num \ code_numeral) <\> term_of_num_semiring (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)" + else if k < 0 then termify (neg_numeral :: num \ int) <\> term_of_num_semiring (2::int) (- k) + else termify (numeral :: num \ int) <\> term_of_num_semiring (2::int) k)" by (simp only: term_of_anything) @@ -201,6 +197,6 @@ hide_const dummy_term valapp -hide_const (open) Const App Abs Free termify valtermify term_of term_of_num_semiring term_of_num_ring tracing +hide_const (open) Const App Abs Free termify valtermify term_of term_of_num_semiring tracing end diff -r 29e92b644d6c -r f067afe98049 src/HOL/Code_Numeral.thy --- a/src/HOL/Code_Numeral.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Code_Numeral.thy Mon Mar 26 10:56:56 2012 +0200 @@ -123,25 +123,6 @@ by (rule equal_refl) -subsection {* Code numerals as datatype of ints *} - -instantiation code_numeral :: number -begin - -definition - "number_of = of_nat o nat" - -instance .. - -end - -lemma nat_of_number [simp]: - "nat_of (number_of k) = number_of k" - by (simp add: number_of_code_numeral_def nat_number_of_def number_of_is_id) - -code_datatype "number_of \ int \ code_numeral" - - subsection {* Basic arithmetic *} instantiation code_numeral :: "{minus, linordered_semidom, semiring_div, linorder}" @@ -176,16 +157,17 @@ end -lemma zero_code_numeral_code [code]: - "(0\code_numeral) = Numeral0" - by (simp add: number_of_code_numeral_def Pls_def) +lemma nat_of_numeral [simp]: "nat_of (numeral k) = numeral k" + by (induct k rule: num_induct) (simp_all add: numeral_inc) -lemma [code_abbrev]: "Numeral0 = (0\code_numeral)" - using zero_code_numeral_code .. +definition Num :: "num \ code_numeral" + where [simp, code_abbrev]: "Num = numeral" + +code_datatype "0::code_numeral" Num lemma one_code_numeral_code [code]: "(1\code_numeral) = Numeral1" - by (simp add: number_of_code_numeral_def Pls_def Bit1_def) + by simp lemma [code_abbrev]: "Numeral1 = (1\code_numeral)" using one_code_numeral_code .. @@ -194,15 +176,8 @@ "of_nat n + of_nat m = of_nat (n + m)" by simp -definition subtract :: "code_numeral \ code_numeral \ code_numeral" where - [simp]: "subtract = minus" - -lemma subtract_code [code nbe]: - "subtract (of_nat n) (of_nat m) = of_nat (n - m)" - by simp - -lemma minus_code_numeral_code [code]: - "minus = subtract" +lemma minus_code_numeral_code [code nbe]: + "of_nat n - of_nat m = of_nat (n - m)" by simp lemma times_code_numeral_code [code nbe]: @@ -281,7 +256,7 @@ qed -hide_const (open) of_nat nat_of Suc subtract int_of +hide_const (open) of_nat nat_of Suc int_of subsection {* Code generator setup *} @@ -298,15 +273,21 @@ (Haskell -) setup {* - Numeral.add_code @{const_name number_code_numeral_inst.number_of_code_numeral} + Numeral.add_code @{const_name Num} false Code_Printer.literal_naive_numeral "SML" - #> fold (Numeral.add_code @{const_name number_code_numeral_inst.number_of_code_numeral} + #> fold (Numeral.add_code @{const_name Num} false Code_Printer.literal_numeral) ["OCaml", "Haskell", "Scala"] *} code_reserved SML Int int code_reserved Eval Integer +code_const "0::code_numeral" + (SML "0") + (OCaml "Big'_int.zero'_big'_int") + (Haskell "0") + (Scala "BigInt(0)") + code_const "plus \ code_numeral \ code_numeral \ code_numeral" (SML "Int.+/ ((_),/ (_))") (OCaml "Big'_int.add'_big'_int") @@ -314,12 +295,12 @@ (Scala infixl 7 "+") (Eval infixl 8 "+") -code_const "Code_Numeral.subtract \ code_numeral \ code_numeral \ code_numeral" - (SML "Int.max/ (_/ -/ _,/ 0 : int)") - (OCaml "Big'_int.max'_big'_int/ (Big'_int.sub'_big'_int/ _/ _)/ Big'_int.zero'_big'_int") - (Haskell "max/ (_/ -/ _)/ (0 :: Integer)") +code_const "minus \ code_numeral \ code_numeral \ code_numeral" + (SML "Int.max/ (0 : int,/ Int.-/ ((_),/ (_)))") + (OCaml "Big'_int.max'_big'_int/ Big'_int.zero'_big'_int/ (Big'_int.sub'_big'_int/ _/ _)") + (Haskell "max/ (0 :: Integer)/ (_/ -/ _)") (Scala "!(_/ -/ _).max(0)") - (Eval "Integer.max/ (_/ -/ _)/ 0") + (Eval "Integer.max/ 0/ (_/ -/ _)") code_const "times \ code_numeral \ code_numeral \ code_numeral" (SML "Int.*/ ((_),/ (_))") diff -r 29e92b644d6c -r f067afe98049 src/HOL/Codegenerator_Test/Generate_Pretty.thy --- a/src/HOL/Codegenerator_Test/Generate_Pretty.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Codegenerator_Test/Generate_Pretty.thy Mon Mar 26 10:56:56 2012 +0200 @@ -10,9 +10,8 @@ lemma [code, code del]: "nat_of_char = nat_of_char" .. lemma [code, code del]: "char_of_nat = char_of_nat" .. -declare Quickcheck_Narrowing.zero_code_int_code[code del] -declare Quickcheck_Narrowing.one_code_int_code[code del] -declare Quickcheck_Narrowing.int_of_code[code del] +declare Quickcheck_Narrowing.one_code_int_code [code del] +declare Quickcheck_Narrowing.int_of_code [code del] subsection {* Check whether generated code compiles *} diff -r 29e92b644d6c -r f067afe98049 src/HOL/Complex.thy --- a/src/HOL/Complex.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Complex.thy Mon Mar 26 10:56:56 2012 +0200 @@ -151,17 +151,6 @@ subsection {* Numerals and Arithmetic *} -instantiation complex :: number_ring -begin - -definition complex_number_of_def: - "number_of w = (of_int w \ complex)" - -instance - by intro_classes (simp only: complex_number_of_def) - -end - lemma complex_Re_of_nat [simp]: "Re (of_nat n) = of_nat n" by (induct n) simp_all @@ -174,14 +163,24 @@ lemma complex_Im_of_int [simp]: "Im (of_int z) = 0" by (cases z rule: int_diff_cases) simp -lemma complex_Re_number_of [simp]: "Re (number_of v) = number_of v" - unfolding number_of_eq by (rule complex_Re_of_int) +lemma complex_Re_numeral [simp]: "Re (numeral v) = numeral v" + using complex_Re_of_int [of "numeral v"] by simp + +lemma complex_Re_neg_numeral [simp]: "Re (neg_numeral v) = neg_numeral v" + using complex_Re_of_int [of "neg_numeral v"] by simp + +lemma complex_Im_numeral [simp]: "Im (numeral v) = 0" + using complex_Im_of_int [of "numeral v"] by simp -lemma complex_Im_number_of [simp]: "Im (number_of v) = 0" - unfolding number_of_eq by (rule complex_Im_of_int) +lemma complex_Im_neg_numeral [simp]: "Im (neg_numeral v) = 0" + using complex_Im_of_int [of "neg_numeral v"] by simp -lemma Complex_eq_number_of [simp]: - "(Complex a b = number_of w) = (a = number_of w \ b = 0)" +lemma Complex_eq_numeral [simp]: + "(Complex a b = numeral w) = (a = numeral w \ b = 0)" + by (simp add: complex_eq_iff) + +lemma Complex_eq_neg_numeral [simp]: + "(Complex a b = neg_numeral w) = (a = neg_numeral w \ b = 0)" by (simp add: complex_eq_iff) @@ -421,7 +420,10 @@ lemma complex_i_not_one [simp]: "ii \ 1" by (simp add: complex_eq_iff) -lemma complex_i_not_number_of [simp]: "ii \ number_of w" +lemma complex_i_not_numeral [simp]: "ii \ numeral w" + by (simp add: complex_eq_iff) + +lemma complex_i_not_neg_numeral [simp]: "ii \ neg_numeral w" by (simp add: complex_eq_iff) lemma i_mult_Complex [simp]: "ii * Complex a b = Complex (- b) a" @@ -505,7 +507,10 @@ lemma complex_cnj_of_int [simp]: "cnj (of_int z) = of_int z" by (simp add: complex_eq_iff) -lemma complex_cnj_number_of [simp]: "cnj (number_of w) = number_of w" +lemma complex_cnj_numeral [simp]: "cnj (numeral w) = numeral w" + by (simp add: complex_eq_iff) + +lemma complex_cnj_neg_numeral [simp]: "cnj (neg_numeral w) = neg_numeral w" by (simp add: complex_eq_iff) lemma complex_cnj_scaleR: "cnj (scaleR r x) = scaleR r (cnj x)" @@ -686,10 +691,10 @@ "(of_nat n :: 'a::linordered_idom) < of_int x \ int n < x" by (metis of_int_of_nat_eq of_int_less_iff) -lemma real_of_nat_less_number_of_iff [simp]: (* TODO: move *) - "real (n::nat) < number_of w \ n < number_of w" - unfolding real_of_nat_def nat_number_of_def number_of_eq - by (simp add: of_nat_less_of_int_iff zless_nat_eq_int_zless) +lemma real_of_nat_less_numeral_iff [simp]: (* TODO: move *) + "real (n::nat) < numeral w \ n < numeral w" + using of_nat_less_of_int_iff [of n "numeral w", where 'a=real] + by (simp add: real_of_nat_def zless_nat_eq_int_zless [symmetric]) lemma arg_unique: assumes "sgn z = cis x" and "-pi < x" and "x \ pi" diff -r 29e92b644d6c -r f067afe98049 src/HOL/Decision_Procs/Approximation.thy --- a/src/HOL/Decision_Procs/Approximation.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Decision_Procs/Approximation.thy Mon Mar 26 10:56:56 2012 +0200 @@ -1350,7 +1350,7 @@ also have "\ \ cos ((?ux - 2 * ?lpi))" using cos_monotone_minus_pi_0'[OF pi_x x_le_ux ux_0] by (simp only: real_of_float_minus real_of_int_minus real_of_one - number_of_Min diff_minus mult_minus_left mult_1_left) + minus_one [symmetric] diff_minus mult_minus_left mult_1_left) also have "\ = cos ((- (?ux - 2 * ?lpi)))" unfolding real_of_float_minus cos_minus .. also have "\ \ (ub_cos prec (- (?ux - 2 * ?lpi)))" @@ -1394,7 +1394,7 @@ also have "\ \ cos ((?lx + 2 * ?lpi))" using cos_monotone_0_pi'[OF lx_0 lx_le_x pi_x] by (simp only: real_of_float_minus real_of_int_minus real_of_one - number_of_Min diff_minus mult_minus_left mult_1_left) + minus_one [symmetric] diff_minus mult_minus_left mult_1_left) also have "\ \ (ub_cos prec (?lx + 2 * ?lpi))" using lb_cos[OF lx_0 pi_lx] by simp finally show ?thesis unfolding u by (simp add: real_of_float_max) @@ -2117,7 +2117,8 @@ lemma interpret_floatarith_num: shows "interpret_floatarith (Num (Float 0 0)) vs = 0" and "interpret_floatarith (Num (Float 1 0)) vs = 1" - and "interpret_floatarith (Num (Float (number_of a) 0)) vs = number_of a" by auto + and "interpret_floatarith (Num (Float (numeral a) 0)) vs = numeral a" + and "interpret_floatarith (Num (Float (neg_numeral a) 0)) vs = neg_numeral a" by auto subsection "Implement approximation function" diff -r 29e92b644d6c -r f067afe98049 src/HOL/Decision_Procs/Cooper.thy --- a/src/HOL/Decision_Procs/Cooper.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Decision_Procs/Cooper.thy Mon Mar 26 10:56:56 2012 +0200 @@ -1883,7 +1883,8 @@ | SOME n => @{code Bound} n) | num_of_term vs @{term "0::int"} = @{code C} 0 | num_of_term vs @{term "1::int"} = @{code C} 1 - | num_of_term vs (@{term "number_of :: int \ int"} $ t) = @{code C} (HOLogic.dest_numeral t) + | num_of_term vs (@{term "numeral :: _ \ int"} $ t) = @{code C} (HOLogic.dest_num t) + | num_of_term vs (@{term "neg_numeral :: _ \ int"} $ t) = @{code C} (~(HOLogic.dest_num t)) | num_of_term vs (Bound i) = @{code Bound} i | num_of_term vs (@{term "uminus :: int \ int"} $ t') = @{code Neg} (num_of_term vs t') | num_of_term vs (@{term "op + :: int \ int \ int"} $ t1 $ t2) = diff -r 29e92b644d6c -r f067afe98049 src/HOL/Decision_Procs/Dense_Linear_Order.thy --- a/src/HOL/Decision_Procs/Dense_Linear_Order.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Decision_Procs/Dense_Linear_Order.thy Mon Mar 26 10:56:56 2012 +0200 @@ -636,14 +636,8 @@ interpretation class_dense_linordered_field: constr_dense_linorder "op <=" "op <" - "\ x y. 1/2 * ((x::'a::{linordered_field,number_ring}) + y)" -proof (unfold_locales, dlo, dlo, auto) - fix x y::'a assume lt: "x < y" - from less_half_sum[OF lt] show "x < (x + y) /2" by simp -next - fix x y::'a assume lt: "x < y" - from gt_half_sum[OF lt] show "(x + y) /2 < y" by simp -qed + "\ x y. 1/2 * ((x::'a::{linordered_field}) + y)" +by (unfold_locales, dlo, dlo, auto) declaration{* let diff -r 29e92b644d6c -r f067afe98049 src/HOL/Decision_Procs/Ferrack.thy --- a/src/HOL/Decision_Procs/Ferrack.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Decision_Procs/Ferrack.thy Mon Mar 26 10:56:56 2012 +0200 @@ -1732,7 +1732,7 @@ (set U \ set U)"using mnz nnz th apply (auto simp add: th add_divide_distrib algebra_simps split_def image_def) by (rule_tac x="(s,m)" in bexI,simp_all) - (rule_tac x="(t,n)" in bexI,simp_all) + (rule_tac x="(t,n)" in bexI,simp_all add: mult_commute) next fix t n s m assume tnU: "(t,n) \ set U" and smU:"(s,m) \ set U" @@ -1937,11 +1937,12 @@ | num_of_term vs (@{term "op * :: real \ real \ real"} $ t1 $ t2) = (case num_of_term vs t1 of @{code C} i => @{code Mul} (i, num_of_term vs t2) | _ => error "num_of_term: unsupported multiplication") - | num_of_term vs (@{term "real :: int \ real"} $ (@{term "number_of :: int \ int"} $ t')) = - @{code C} (HOLogic.dest_numeral t') - | num_of_term vs (@{term "number_of :: int \ real"} $ t') = - @{code C} (HOLogic.dest_numeral t') - | num_of_term vs t = error ("num_of_term: unknown term"); + | num_of_term vs (@{term "real :: int \ real"} $ t') = + (@{code C} (snd (HOLogic.dest_number t')) + handle TERM _ => error ("num_of_term: unknown term")) + | num_of_term vs t' = + (@{code C} (snd (HOLogic.dest_number t')) + handle TERM _ => error ("num_of_term: unknown term")); fun fm_of_term vs @{term True} = @{code T} | fm_of_term vs @{term False} = @{code F} diff -r 29e92b644d6c -r f067afe98049 src/HOL/Decision_Procs/MIR.thy --- a/src/HOL/Decision_Procs/MIR.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Decision_Procs/MIR.thy Mon Mar 26 10:56:56 2012 +0200 @@ -4901,7 +4901,7 @@ (set U \ set U)"using mnz nnz th apply (auto simp add: th add_divide_distrib algebra_simps split_def image_def) by (rule_tac x="(s,m)" in bexI,simp_all) - (rule_tac x="(t,n)" in bexI,simp_all) + (rule_tac x="(t,n)" in bexI,simp_all add: mult_commute) next fix t n s m assume tnU: "(t,n) \ set U" and smU:"(s,m) \ set U" @@ -5536,14 +5536,18 @@ (case (num_of_term vs t1) of @{code C} i => @{code Mul} (i, num_of_term vs t2) | _ => error "num_of_term: unsupported Multiplication") - | num_of_term vs (@{term "real :: int \ real"} $ (@{term "number_of :: int \ int"} $ t')) = - @{code C} (HOLogic.dest_numeral t') + | num_of_term vs (@{term "real :: int \ real"} $ (@{term "numeral :: _ \ int"} $ t')) = + @{code C} (HOLogic.dest_num t') + | num_of_term vs (@{term "real :: int \ real"} $ (@{term "neg_numeral :: _ \ int"} $ t')) = + @{code C} (~ (HOLogic.dest_num t')) | num_of_term vs (@{term "real :: int \ real"} $ (@{term "floor :: real \ int"} $ t')) = @{code Floor} (num_of_term vs t') | num_of_term vs (@{term "real :: int \ real"} $ (@{term "ceiling :: real \ int"} $ t')) = @{code Neg} (@{code Floor} (@{code Neg} (num_of_term vs t'))) - | num_of_term vs (@{term "number_of :: int \ real"} $ t') = - @{code C} (HOLogic.dest_numeral t') + | num_of_term vs (@{term "numeral :: _ \ real"} $ t') = + @{code C} (HOLogic.dest_num t') + | num_of_term vs (@{term "neg_numeral :: _ \ real"} $ t') = + @{code C} (~ (HOLogic.dest_num t')) | num_of_term vs t = error ("num_of_term: unknown term " ^ Syntax.string_of_term @{context} t); fun fm_of_term vs @{term True} = @{code T} @@ -5554,8 +5558,10 @@ @{code Le} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) | fm_of_term vs (@{term "op = :: real \ real \ bool"} $ t1 $ t2) = @{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) - | fm_of_term vs (@{term "op rdvd"} $ (@{term "real :: int \ real"} $ (@{term "number_of :: int \ int"} $ t1)) $ t2) = - @{code Dvd} (HOLogic.dest_numeral t1, num_of_term vs t2) + | fm_of_term vs (@{term "op rdvd"} $ (@{term "real :: int \ real"} $ (@{term "numeral :: _ \ int"} $ t1)) $ t2) = + @{code Dvd} (HOLogic.dest_num t1, num_of_term vs t2) + | fm_of_term vs (@{term "op rdvd"} $ (@{term "real :: int \ real"} $ (@{term "neg_numeral :: _ \ int"} $ t1)) $ t2) = + @{code Dvd} (~ (HOLogic.dest_num t1), num_of_term vs t2) | fm_of_term vs (@{term "op = :: bool \ bool \ bool"} $ t1 $ t2) = @{code Iff} (fm_of_term vs t1, fm_of_term vs t2) | fm_of_term vs (@{term HOL.conj} $ t1 $ t2) = diff -r 29e92b644d6c -r f067afe98049 src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy --- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Mon Mar 26 10:56:56 2012 +0200 @@ -25,7 +25,7 @@ | "tmsize (CNP n c a) = 3 + polysize c + tmsize a " (* Semantics of terms tm *) -primrec Itm :: "'a::{field_char_0, field_inverse_zero, number_ring} list \ 'a list \ tm \ 'a" where +primrec Itm :: "'a::{field_char_0, field_inverse_zero} list \ 'a list \ tm \ 'a" where "Itm vs bs (CP c) = (Ipoly vs c)" | "Itm vs bs (Bound n) = bs!n" | "Itm vs bs (Neg a) = -(Itm vs bs a)" @@ -430,7 +430,7 @@ by (induct p rule: fmsize.induct) simp_all (* Semantics of formulae (fm) *) -primrec Ifm ::"'a::{linordered_field_inverse_zero, number_ring} list \ 'a list \ fm \ bool" where +primrec Ifm ::"'a::{linordered_field_inverse_zero} list \ 'a list \ fm \ bool" where "Ifm vs bs T = True" | "Ifm vs bs F = False" | "Ifm vs bs (Lt a) = (Itm vs bs a < 0)" @@ -1937,7 +1937,7 @@ also have "\ \ - (?a * ?s) + 2*?d*?r = 0" using d by simp finally have ?thesis using c d - by (simp add: r[of "- (Itm vs (x # bs) s / (2 * \d\\<^sub>p\<^bsup>vs\<^esup>))"] msubsteq_def Let_def evaldjf_ex del: one_add_one_is_two) + by (simp add: r[of "- (Itm vs (x # bs) s / (2 * \d\\<^sub>p\<^bsup>vs\<^esup>))"] msubsteq_def Let_def evaldjf_ex) } moreover {assume c: "?c \ 0" and d: "?d=0" @@ -1950,7 +1950,7 @@ by (simp add: field_simps right_distrib[of "2*?c"] del: right_distrib) also have "\ \ - (?a * ?t) + 2*?c*?r = 0" using c by simp finally have ?thesis using c d - by (simp add: r[of "- (?t/ (2*?c))"] msubsteq_def Let_def evaldjf_ex del: one_add_one_is_two) + by (simp add: r[of "- (?t/ (2*?c))"] msubsteq_def Let_def evaldjf_ex) } moreover {assume c: "?c \ 0" and d: "?d\0" hence dc: "?c * ?d *2 \ 0" by simp @@ -2019,7 +2019,7 @@ also have "\ \ - (?a * ?s) + 2*?d*?r \ 0" using d by simp finally have ?thesis using c d - by (simp add: r[of "- (Itm vs (x # bs) s / (2 * \d\\<^sub>p\<^bsup>vs\<^esup>))"] msubstneq_def Let_def evaldjf_ex del: one_add_one_is_two) + by (simp add: r[of "- (Itm vs (x # bs) s / (2 * \d\\<^sub>p\<^bsup>vs\<^esup>))"] msubstneq_def Let_def evaldjf_ex) } moreover {assume c: "?c \ 0" and d: "?d=0" @@ -2032,7 +2032,7 @@ by (simp add: field_simps right_distrib[of "2*?c"] del: right_distrib) also have "\ \ - (?a * ?t) + 2*?c*?r \ 0" using c by simp finally have ?thesis using c d - by (simp add: r[of "- (?t/ (2*?c))"] msubstneq_def Let_def evaldjf_ex del: one_add_one_is_two) + by (simp add: r[of "- (?t/ (2*?c))"] msubstneq_def Let_def evaldjf_ex) } moreover {assume c: "?c \ 0" and d: "?d\0" hence dc: "?c * ?d *2 \ 0" by simp @@ -2616,10 +2616,10 @@ using lp tnb by (simp add: msubst2_def msubstneg_nb msubstpos_nb conj_nb disj_nb lt_nb simpfm_bound0) -lemma mult_minus2_left: "-2 * (x::'a::number_ring) = - (2 * x)" +lemma mult_minus2_left: "-2 * (x::'a::comm_ring_1) = - (2 * x)" by simp -lemma mult_minus2_right: "(x::'a::number_ring) * -2 = - (x * 2)" +lemma mult_minus2_right: "(x::'a::comm_ring_1) * -2 = - (x * 2)" by simp lemma islin_qf: "islin p \ qfree p" @@ -3005,11 +3005,11 @@ *} "parametric QE for linear Arithmetic over fields, Version 2" -lemma "\(x::'a::{linordered_field_inverse_zero, number_ring}). y \ -1 \ (y + 1)*x < 0" - apply (frpar type: "'a::{linordered_field_inverse_zero, number_ring}" pars: "y::'a::{linordered_field_inverse_zero, number_ring}") +lemma "\(x::'a::{linordered_field_inverse_zero}). y \ -1 \ (y + 1)*x < 0" + apply (frpar type: "'a::{linordered_field_inverse_zero}" pars: "y::'a::{linordered_field_inverse_zero}") apply (simp add: field_simps) apply (rule spec[where x=y]) - apply (frpar type: "'a::{linordered_field_inverse_zero, number_ring}" pars: "z::'a::{linordered_field_inverse_zero, number_ring}") + apply (frpar type: "'a::{linordered_field_inverse_zero}" pars: "z::'a::{linordered_field_inverse_zero}") by simp text{* Collins/Jones Problem *} @@ -3030,11 +3030,11 @@ oops *) -lemma "\(x::'a::{linordered_field_inverse_zero, number_ring}). y \ -1 \ (y + 1)*x < 0" - apply (frpar2 type: "'a::{linordered_field_inverse_zero, number_ring}" pars: "y::'a::{linordered_field_inverse_zero, number_ring}") +lemma "\(x::'a::{linordered_field_inverse_zero}). y \ -1 \ (y + 1)*x < 0" + apply (frpar2 type: "'a::{linordered_field_inverse_zero}" pars: "y::'a::{linordered_field_inverse_zero}") apply (simp add: field_simps) apply (rule spec[where x=y]) - apply (frpar2 type: "'a::{linordered_field_inverse_zero, number_ring}" pars: "z::'a::{linordered_field_inverse_zero, number_ring}") + apply (frpar2 type: "'a::{linordered_field_inverse_zero}" pars: "z::'a::{linordered_field_inverse_zero}") by simp text{* Collins/Jones Problem *} diff -r 29e92b644d6c -r f067afe98049 src/HOL/Decision_Procs/cooper_tac.ML --- a/src/HOL/Decision_Procs/cooper_tac.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Decision_Procs/cooper_tac.ML Mon Mar 26 10:56:56 2012 +0200 @@ -18,15 +18,12 @@ val cooper_ss = @{simpset}; val nT = HOLogic.natT; -val binarith = @{thms normalize_bin_simps}; -val comp_arith = binarith @ @{thms simp_thms}; +val comp_arith = @{thms simp_thms} val zdvd_int = @{thm zdvd_int}; val zdiff_int_split = @{thm zdiff_int_split}; val all_nat = @{thm all_nat}; val ex_nat = @{thm ex_nat}; -val number_of1 = @{thm number_of1}; -val number_of2 = @{thm number_of2}; val split_zdiv = @{thm split_zdiv}; val split_zmod = @{thm split_zmod}; val mod_div_equality' = @{thm mod_div_equality'}; @@ -90,14 +87,13 @@ [split_zdiv, split_zmod, split_div', @{thm "split_min"}, @{thm "split_max"}] (* Simp rules for changing (n::int) to int n *) val simpset1 = HOL_basic_ss - addsimps [@{thm nat_number_of_def}, zdvd_int] @ map (fn r => r RS sym) - [@{thm int_int_eq}, @{thm zle_int}, @{thm zless_int}, @{thm zadd_int}, @{thm zmult_int}] + addsimps [zdvd_int] @ map (fn r => r RS sym) + [@{thm int_numeral}, @{thm int_int_eq}, @{thm zle_int}, @{thm zless_int}, @{thm zadd_int}, @{thm zmult_int}] |> Splitter.add_split zdiff_int_split (*simp rules for elimination of int n*) val simpset2 = HOL_basic_ss - addsimps [@{thm nat_0_le}, @{thm all_nat}, @{thm ex_nat}, - @{thm number_of1}, @{thm number_of2}, @{thm int_0}, @{thm int_1}] + addsimps [@{thm nat_0_le}, @{thm all_nat}, @{thm ex_nat}, @{thm zero_le_numeral}, @{thm order_refl}(* FIXME: necessary? *), @{thm int_0}, @{thm int_1}] |> fold Simplifier.add_cong [@{thm conj_le_cong}, @{thm imp_le_cong}] (* simp rules for elimination of abs *) val simpset3 = HOL_basic_ss |> Splitter.add_split @{thm abs_split} diff -r 29e92b644d6c -r f067afe98049 src/HOL/Decision_Procs/ex/Dense_Linear_Order_Ex.thy --- a/src/HOL/Decision_Procs/ex/Dense_Linear_Order_Ex.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Decision_Procs/ex/Dense_Linear_Order_Ex.thy Mon Mar 26 10:56:56 2012 +0200 @@ -7,147 +7,147 @@ begin lemma - "\(y::'a::{linordered_field_inverse_zero, number_ring}) <2. x + 3* y < 0 \ x - y >0" + "\(y::'a::{linordered_field_inverse_zero}) <2. x + 3* y < 0 \ x - y >0" by ferrack -lemma "~ (ALL x (y::'a::{linordered_field_inverse_zero, number_ring}). x < y --> 10*x < 11*y)" +lemma "~ (ALL x (y::'a::{linordered_field_inverse_zero}). x < y --> 10*x < 11*y)" by ferrack -lemma "ALL (x::'a::{linordered_field_inverse_zero, number_ring}) y. x < y --> (10*(x + 5*y + -1) < 60*y)" +lemma "ALL (x::'a::{linordered_field_inverse_zero}) y. x < y --> (10*(x + 5*y + -1) < 60*y)" by ferrack -lemma "EX (x::'a::{linordered_field_inverse_zero, number_ring}) y. x ~= y --> x < y" +lemma "EX (x::'a::{linordered_field_inverse_zero}) y. x ~= y --> x < y" by ferrack -lemma "EX (x::'a::{linordered_field_inverse_zero, number_ring}) y. (x ~= y & 10*x ~= 9*y & 10*x < y) --> x < y" +lemma "EX (x::'a::{linordered_field_inverse_zero}) y. (x ~= y & 10*x ~= 9*y & 10*x < y) --> x < y" by ferrack -lemma "ALL (x::'a::{linordered_field_inverse_zero, number_ring}) y. (x ~= y & 5*x <= y) --> 500*x <= 100*y" +lemma "ALL (x::'a::{linordered_field_inverse_zero}) y. (x ~= y & 5*x <= y) --> 500*x <= 100*y" by ferrack -lemma "ALL (x::'a::{linordered_field_inverse_zero, number_ring}). (EX (y::'a::{linordered_field_inverse_zero, number_ring}). 4*x + 3*y <= 0 & 4*x + 3*y >= -1)" +lemma "ALL (x::'a::{linordered_field_inverse_zero}). (EX (y::'a::{linordered_field_inverse_zero}). 4*x + 3*y <= 0 & 4*x + 3*y >= -1)" by ferrack -lemma "ALL (x::'a::{linordered_field_inverse_zero, number_ring}) < 0. (EX (y::'a::{linordered_field_inverse_zero, number_ring}) > 0. 7*x + y > 0 & x - y <= 9)" +lemma "ALL (x::'a::{linordered_field_inverse_zero}) < 0. (EX (y::'a::{linordered_field_inverse_zero}) > 0. 7*x + y > 0 & x - y <= 9)" by ferrack -lemma "EX (x::'a::{linordered_field_inverse_zero, number_ring}). (0 < x & x < 1) --> (ALL y > 1. x + y ~= 1)" +lemma "EX (x::'a::{linordered_field_inverse_zero}). (0 < x & x < 1) --> (ALL y > 1. x + y ~= 1)" by ferrack -lemma "EX x. (ALL (y::'a::{linordered_field_inverse_zero, number_ring}). y < 2 --> 2*(y - x) \ 0 )" +lemma "EX x. (ALL (y::'a::{linordered_field_inverse_zero}). y < 2 --> 2*(y - x) \ 0 )" by ferrack -lemma "ALL (x::'a::{linordered_field_inverse_zero, number_ring}). x < 10 | x > 20 | (EX y. y>= 0 & y <= 10 & x+y = 20)" +lemma "ALL (x::'a::{linordered_field_inverse_zero}). x < 10 | x > 20 | (EX y. y>= 0 & y <= 10 & x+y = 20)" by ferrack -lemma "ALL (x::'a::{linordered_field_inverse_zero, number_ring}) y z. x + y < z --> y >= z --> x < 0" +lemma "ALL (x::'a::{linordered_field_inverse_zero}) y z. x + y < z --> y >= z --> x < 0" by ferrack -lemma "EX (x::'a::{linordered_field_inverse_zero, number_ring}) y z. x + 7*y < 5* z & 5*y >= 7*z & x < 0" +lemma "EX (x::'a::{linordered_field_inverse_zero}) y z. x + 7*y < 5* z & 5*y >= 7*z & x < 0" by ferrack -lemma "ALL (x::'a::{linordered_field_inverse_zero, number_ring}) y z. abs (x + y) <= z --> (abs z = z)" +lemma "ALL (x::'a::{linordered_field_inverse_zero}) y z. abs (x + y) <= z --> (abs z = z)" by ferrack -lemma "EX (x::'a::{linordered_field_inverse_zero, number_ring}) y z. x + 7*y - 5* z < 0 & 5*y + 7*z + 3*x < 0" +lemma "EX (x::'a::{linordered_field_inverse_zero}) y z. x + 7*y - 5* z < 0 & 5*y + 7*z + 3*x < 0" by ferrack -lemma "ALL (x::'a::{linordered_field_inverse_zero, number_ring}) y z. (abs (5*x+3*y+z) <= 5*x+3*y+z & abs (5*x+3*y+z) >= - (5*x+3*y+z)) | (abs (5*x+3*y+z) >= 5*x+3*y+z & abs (5*x+3*y+z) <= - (5*x+3*y+z))" +lemma "ALL (x::'a::{linordered_field_inverse_zero}) y z. (abs (5*x+3*y+z) <= 5*x+3*y+z & abs (5*x+3*y+z) >= - (5*x+3*y+z)) | (abs (5*x+3*y+z) >= 5*x+3*y+z & abs (5*x+3*y+z) <= - (5*x+3*y+z))" by ferrack -lemma "ALL (x::'a::{linordered_field_inverse_zero, number_ring}) y. x < y --> (EX z>0. x+z = y)" +lemma "ALL (x::'a::{linordered_field_inverse_zero}) y. x < y --> (EX z>0. x+z = y)" by ferrack -lemma "ALL (x::'a::{linordered_field_inverse_zero, number_ring}) y. x < y --> (EX z>0. x+z = y)" +lemma "ALL (x::'a::{linordered_field_inverse_zero}) y. x < y --> (EX z>0. x+z = y)" by ferrack -lemma "ALL (x::'a::{linordered_field_inverse_zero, number_ring}) y. (EX z>0. abs (x - y) <= z )" +lemma "ALL (x::'a::{linordered_field_inverse_zero}) y. (EX z>0. abs (x - y) <= z )" by ferrack -lemma "EX (x::'a::{linordered_field_inverse_zero, number_ring}) y. (ALL z<0. (z < x --> z <= y) & (z > y --> z >= x))" +lemma "EX (x::'a::{linordered_field_inverse_zero}) y. (ALL z<0. (z < x --> z <= y) & (z > y --> z >= x))" by ferrack -lemma "EX (x::'a::{linordered_field_inverse_zero, number_ring}) y. (ALL z>=0. abs (3*x+7*y) <= 2*z + 1)" +lemma "EX (x::'a::{linordered_field_inverse_zero}) y. (ALL z>=0. abs (3*x+7*y) <= 2*z + 1)" by ferrack -lemma "EX (x::'a::{linordered_field_inverse_zero, number_ring}) y. (ALL z<0. (z < x --> z <= y) & (z > y --> z >= x))" +lemma "EX (x::'a::{linordered_field_inverse_zero}) y. (ALL z<0. (z < x --> z <= y) & (z > y --> z >= x))" by ferrack -lemma "EX (x::'a::{linordered_field_inverse_zero, number_ring})>0. (ALL y. (EX z. 13* abs z \ abs (12*y - x) & 5*x - 3*(abs y) <= 7*z))" +lemma "EX (x::'a::{linordered_field_inverse_zero})>0. (ALL y. (EX z. 13* abs z \ abs (12*y - x) & 5*x - 3*(abs y) <= 7*z))" by ferrack -lemma "EX (x::'a::{linordered_field_inverse_zero, number_ring}). abs (4*x + 17) < 4 & (ALL y . abs (x*34 - 34*y - 9) \ 0 \ (EX z. 5*x - 3*abs y <= 7*z))" +lemma "EX (x::'a::{linordered_field_inverse_zero}). abs (4*x + 17) < 4 & (ALL y . abs (x*34 - 34*y - 9) \ 0 \ (EX z. 5*x - 3*abs y <= 7*z))" by ferrack -lemma "ALL (x::'a::{linordered_field_inverse_zero, number_ring}). (EX y > abs (23*x - 9). (ALL z > abs (3*y - 19* abs x). x+z > 2*y))" +lemma "ALL (x::'a::{linordered_field_inverse_zero}). (EX y > abs (23*x - 9). (ALL z > abs (3*y - 19* abs x). x+z > 2*y))" by ferrack -lemma "ALL (x::'a::{linordered_field_inverse_zero, number_ring}). (EX y< abs (3*x - 1). (ALL z >= (3*abs x - 1). abs (12*x - 13*y + 19*z) > abs (23*x) ))" +lemma "ALL (x::'a::{linordered_field_inverse_zero}). (EX y< abs (3*x - 1). (ALL z >= (3*abs x - 1). abs (12*x - 13*y + 19*z) > abs (23*x) ))" by ferrack -lemma "EX (x::'a::{linordered_field_inverse_zero, number_ring}). abs x < 100 & (ALL y > x. (EX z<2*y - x. 5*x - 3*y <= 7*z))" +lemma "EX (x::'a::{linordered_field_inverse_zero}). abs x < 100 & (ALL y > x. (EX z<2*y - x. 5*x - 3*y <= 7*z))" by ferrack -lemma "ALL (x::'a::{linordered_field_inverse_zero, number_ring}) y z w. 7*x<3*y --> 5*y < 7*z --> z < 2*w --> 7*(2*w-x) > 2*y" +lemma "ALL (x::'a::{linordered_field_inverse_zero}) y z w. 7*x<3*y --> 5*y < 7*z --> z < 2*w --> 7*(2*w-x) > 2*y" by ferrack -lemma "EX (x::'a::{linordered_field_inverse_zero, number_ring}) y z w. 5*x + 3*z - 17*w + abs (y - 8*x + z) <= 89" +lemma "EX (x::'a::{linordered_field_inverse_zero}) y z w. 5*x + 3*z - 17*w + abs (y - 8*x + z) <= 89" by ferrack -lemma "EX (x::'a::{linordered_field_inverse_zero, number_ring}) y z w. 5*x + 3*z - 17*w + 7* (y - 8*x + z) <= max y (7*z - x + w)" +lemma "EX (x::'a::{linordered_field_inverse_zero}) y z w. 5*x + 3*z - 17*w + 7* (y - 8*x + z) <= max y (7*z - x + w)" by ferrack -lemma "EX (x::'a::{linordered_field_inverse_zero, number_ring}) y z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w)" +lemma "EX (x::'a::{linordered_field_inverse_zero}) y z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w)" by ferrack -lemma "ALL (x::'a::{linordered_field_inverse_zero, number_ring}) y z. (EX w >= (x+y+z). w <= abs x + abs y + abs z)" +lemma "ALL (x::'a::{linordered_field_inverse_zero}) y z. (EX w >= (x+y+z). w <= abs x + abs y + abs z)" by ferrack -lemma "~(ALL (x::'a::{linordered_field_inverse_zero, number_ring}). (EX y z w. 3* x + z*4 = 3*y & x + y < z & x> w & 3*x < w + y))" +lemma "~(ALL (x::'a::{linordered_field_inverse_zero}). (EX y z w. 3* x + z*4 = 3*y & x + y < z & x> w & 3*x < w + y))" by ferrack -lemma "ALL (x::'a::{linordered_field_inverse_zero, number_ring}) y. (EX z w. abs (x-y) = (z-w) & z*1234 < 233*x & w ~= y)" +lemma "ALL (x::'a::{linordered_field_inverse_zero}) y. (EX z w. abs (x-y) = (z-w) & z*1234 < 233*x & w ~= y)" by ferrack -lemma "ALL (x::'a::{linordered_field_inverse_zero, number_ring}). (EX y z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w))" +lemma "ALL (x::'a::{linordered_field_inverse_zero}). (EX y z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w))" by ferrack -lemma "EX (x::'a::{linordered_field_inverse_zero, number_ring}) y z. (ALL w >= abs (x+y+z). w >= abs x + abs y + abs z)" +lemma "EX (x::'a::{linordered_field_inverse_zero}) y z. (ALL w >= abs (x+y+z). w >= abs x + abs y + abs z)" by ferrack -lemma "EX z. (ALL (x::'a::{linordered_field_inverse_zero, number_ring}) y. (EX w >= (x+y+z). w <= abs x + abs y + abs z))" +lemma "EX z. (ALL (x::'a::{linordered_field_inverse_zero}) y. (EX w >= (x+y+z). w <= abs x + abs y + abs z))" by ferrack -lemma "EX z. (ALL (x::'a::{linordered_field_inverse_zero, number_ring}) < abs z. (EX y w. x< y & x < z & x> w & 3*x < w + y))" +lemma "EX z. (ALL (x::'a::{linordered_field_inverse_zero}) < abs z. (EX y w. x< y & x < z & x> w & 3*x < w + y))" by ferrack -lemma "ALL (x::'a::{linordered_field_inverse_zero, number_ring}) y. (EX z. (ALL w. abs (x-y) = abs (z-w) --> z < x & w ~= y))" +lemma "ALL (x::'a::{linordered_field_inverse_zero}) y. (EX z. (ALL w. abs (x-y) = abs (z-w) --> z < x & w ~= y))" by ferrack -lemma "EX y. (ALL (x::'a::{linordered_field_inverse_zero, number_ring}). (EX z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w)))" +lemma "EX y. (ALL (x::'a::{linordered_field_inverse_zero}). (EX z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w)))" by ferrack -lemma "EX (x::'a::{linordered_field_inverse_zero, number_ring}) z. (ALL w >= 13*x - 4*z. (EX y. w >= abs x + abs y + z))" +lemma "EX (x::'a::{linordered_field_inverse_zero}) z. (ALL w >= 13*x - 4*z. (EX y. w >= abs x + abs y + z))" by ferrack -lemma "EX (x::'a::{linordered_field_inverse_zero, number_ring}). (ALL y < x. (EX z > (x+y). +lemma "EX (x::'a::{linordered_field_inverse_zero}). (ALL y < x. (EX z > (x+y). (ALL w. 5*w + 10*x - z >= y --> w + 7*x + 3*z >= 2*y)))" by ferrack -lemma "EX (x::'a::{linordered_field_inverse_zero, number_ring}). (ALL y. (EX z > y. +lemma "EX (x::'a::{linordered_field_inverse_zero}). (ALL y. (EX z > y. (ALL w . w < 13 --> w + 10*x - z >= y --> 5*w + 7*x + 13*z >= 2*y)))" by ferrack -lemma "EX (x::'a::{linordered_field_inverse_zero, number_ring}) y z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w)" +lemma "EX (x::'a::{linordered_field_inverse_zero}) y z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w)" by ferrack -lemma "ALL (x::'a::{linordered_field_inverse_zero, number_ring}). (EX y. (ALL z>19. y <= x + z & (EX w. abs (y - x) < w)))" +lemma "ALL (x::'a::{linordered_field_inverse_zero}). (EX y. (ALL z>19. y <= x + z & (EX w. abs (y - x) < w)))" by ferrack -lemma "ALL (x::'a::{linordered_field_inverse_zero, number_ring}). (EX y. (ALL z>19. y <= x + z & (EX w. abs (x + z) < w - y)))" +lemma "ALL (x::'a::{linordered_field_inverse_zero}). (EX y. (ALL z>19. y <= x + z & (EX w. abs (x + z) < w - y)))" by ferrack -lemma "ALL (x::'a::{linordered_field_inverse_zero, number_ring}). (EX y. abs y ~= abs x & (ALL z> max x y. (EX w. w ~= y & w ~= z & 3*w - z >= x + y)))" +lemma "ALL (x::'a::{linordered_field_inverse_zero}). (EX y. abs y ~= abs x & (ALL z> max x y. (EX w. w ~= y & w ~= z & 3*w - z >= x + y)))" by ferrack end diff -r 29e92b644d6c -r f067afe98049 src/HOL/Decision_Procs/ferrack_tac.ML --- a/src/HOL/Decision_Procs/ferrack_tac.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Decision_Procs/ferrack_tac.ML Mon Mar 26 10:56:56 2012 +0200 @@ -20,17 +20,13 @@ in @{simpset} delsimps ths addsimps (map (fn th => th RS sym) ths) end; -val binarith = - @{thms normalize_bin_simps} @ @{thms pred_bin_simps} @ @{thms succ_bin_simps} @ - @{thms add_bin_simps} @ @{thms minus_bin_simps} @ @{thms mult_bin_simps}; -val comp_arith = binarith @ @{thms simp_thms}; +val binarith = @{thms arith_simps} +val comp_arith = binarith @ @{thms simp_thms} val zdvd_int = @{thm zdvd_int}; val zdiff_int_split = @{thm zdiff_int_split}; val all_nat = @{thm all_nat}; val ex_nat = @{thm ex_nat}; -val number_of1 = @{thm number_of1}; -val number_of2 = @{thm number_of2}; val split_zdiv = @{thm split_zdiv}; val split_zmod = @{thm split_zmod}; val mod_div_equality' = @{thm mod_div_equality'}; diff -r 29e92b644d6c -r f067afe98049 src/HOL/Decision_Procs/mir_tac.ML --- a/src/HOL/Decision_Procs/mir_tac.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Decision_Procs/mir_tac.ML Mon Mar 26 10:56:56 2012 +0200 @@ -21,16 +21,15 @@ end; val nT = HOLogic.natT; - val nat_arith = [@{thm "add_nat_number_of"}, @{thm "diff_nat_number_of"}, - @{thm "mult_nat_number_of"}, @{thm "eq_nat_number_of"}, @{thm "less_nat_number_of"}]; + val nat_arith = [@{thm diff_nat_numeral}]; val comp_arith = [@{thm "Let_def"}, @{thm "if_False"}, @{thm "if_True"}, @{thm "add_0"}, - @{thm "add_Suc"}, @{thm "add_number_of_left"}, @{thm "mult_number_of_left"}, + @{thm "add_Suc"}, @{thm add_numeral_left}, @{thm mult_numeral_left(1)}, @{thm "Suc_eq_plus1"}] @ - (map (fn th => th RS sym) [@{thm "numeral_1_eq_1"}, @{thm "numeral_0_eq_0"}]) + (map (fn th => th RS sym) [@{thm "numeral_1_eq_1"}]) @ @{thms arith_simps} @ nat_arith @ @{thms rel_simps} val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_right"}, - @{thm "real_of_nat_number_of"}, + @{thm real_of_nat_numeral}, @{thm "real_of_nat_Suc"}, @{thm "real_of_nat_one"}, @{thm "real_of_one"}, @{thm "real_of_int_zero"}, @{thm "real_of_nat_zero"}, @{thm "divide_zero"}, @@ -44,8 +43,6 @@ val zdiff_int_split = @{thm "zdiff_int_split"}; val all_nat = @{thm "all_nat"}; val ex_nat = @{thm "ex_nat"}; -val number_of1 = @{thm "number_of1"}; -val number_of2 = @{thm "number_of2"}; val split_zdiv = @{thm "split_zdiv"}; val split_zmod = @{thm "split_zmod"}; val mod_div_equality' = @{thm "mod_div_equality'"}; @@ -113,15 +110,15 @@ @{thm "split_min"}, @{thm "split_max"}] (* Simp rules for changing (n::int) to int n *) val simpset1 = HOL_basic_ss - addsimps [@{thm "nat_number_of_def"}, @{thm "zdvd_int"}] @ map (fn r => r RS sym) + addsimps [@{thm "zdvd_int"}] @ map (fn r => r RS sym) [@{thm "int_int_eq"}, @{thm "zle_int"}, @{thm "zless_int"}, @{thm "zadd_int"}, - @{thm "zmult_int"}] + @{thm nat_numeral}, @{thm "zmult_int"}] |> Splitter.add_split @{thm "zdiff_int_split"} (*simp rules for elimination of int n*) val simpset2 = HOL_basic_ss - addsimps [@{thm "nat_0_le"}, @{thm "all_nat"}, @{thm "ex_nat"}, @{thm "number_of1"}, - @{thm "number_of2"}, @{thm "int_0"}, @{thm "int_1"}] + addsimps [@{thm "nat_0_le"}, @{thm "all_nat"}, @{thm "ex_nat"}, @{thm zero_le_numeral}, + @{thm "int_0"}, @{thm "int_1"}] |> fold Simplifier.add_cong [@{thm "conj_le_cong"}, @{thm "imp_le_cong"}] (* simp rules for elimination of abs *) val ct = cterm_of thy (HOLogic.mk_Trueprop t) diff -r 29e92b644d6c -r f067afe98049 src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Deriv.thy Mon Mar 26 10:56:56 2012 +0200 @@ -186,7 +186,6 @@ apply (erule DERIV_mult') apply (erule (1) DERIV_inverse') apply (simp add: ring_distribs nonzero_inverse_mult_distrib) -apply (simp add: mult_ac) done lemma DERIV_power_Suc: diff -r 29e92b644d6c -r f067afe98049 src/HOL/Divides.thy --- a/src/HOL/Divides.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Divides.thy Mon Mar 26 10:56:56 2012 +0200 @@ -1138,8 +1138,8 @@ lemma Suc_mod_eq_add3_mod: "(Suc (Suc (Suc m))) mod n = (3+m) mod n" by (simp add: Suc3_eq_add_3) -lemmas Suc_div_eq_add3_div_number_of [simp] = Suc_div_eq_add3_div [of _ "number_of v"] for v -lemmas Suc_mod_eq_add3_mod_number_of [simp] = Suc_mod_eq_add3_mod [of _ "number_of v"] for v +lemmas Suc_div_eq_add3_div_numeral [simp] = Suc_div_eq_add3_div [of _ "numeral v"] for v +lemmas Suc_mod_eq_add3_mod_numeral [simp] = Suc_mod_eq_add3_mod [of _ "numeral v"] for v lemma Suc_times_mod_eq: "1 Suc (k * m) mod k = 1" @@ -1147,7 +1147,7 @@ apply (simp_all add: mod_Suc) done -declare Suc_times_mod_eq [of "number_of w", simp] for w +declare Suc_times_mod_eq [of "numeral w", simp] for w lemma [simp]: "n div k \ (Suc n) div k" by (simp add: div_le_mono) @@ -1177,17 +1177,22 @@ apply (subst mod_Suc [of "m mod n"], simp) done +lemma mod_2_not_eq_zero_eq_one_nat: + fixes n :: nat + shows "n mod 2 \ 0 \ n mod 2 = 1" + by simp + subsection {* Division on @{typ int} *} definition divmod_int_rel :: "int \ int \ int \ int \ bool" where --{*definition of quotient and remainder*} - [code]: "divmod_int_rel a b = (\(q, r). a = b * q + r \ + "divmod_int_rel a b = (\(q, r). a = b * q + r \ (if 0 < b then 0 \ r \ r < b else b < r \ r \ 0))" definition adjust :: "int \ int \ int \ int \ int" where --{*for the division algorithm*} - [code]: "adjust b = (\(q, r). if 0 \ r - b then (2 * q + 1, r - b) + "adjust b = (\(q, r). if 0 \ r - b then (2 * q + 1, r - b) else (2 * q, r))" text{*algorithm for the case @{text "a\0, b>0"}*} @@ -1318,11 +1323,11 @@ text{*And positive divisors*} lemma adjust_eq [simp]: - "adjust b (q,r) = - (let diff = r-b in - if 0 \ diff then (2*q + 1, diff) + "adjust b (q, r) = + (let diff = r - b in + if 0 \ diff then (2 * q + 1, diff) else (2*q, r))" -by (simp add: Let_def adjust_def) + by (simp add: Let_def adjust_def) declare posDivAlg.simps [simp del] @@ -1420,6 +1425,9 @@ text {* Tool setup *} +(* FIXME: Theorem list add_0s doesn't exist, because Numeral0 has gone. *) +lemmas add_0s = add_0_left add_0_right + ML {* structure Cancel_Div_Mod_Int = Cancel_Div_Mod ( @@ -1674,16 +1682,6 @@ by (rule divmod_int_rel_mod [of a b q r], simp add: divmod_int_rel_def) -lemmas arithmetic_simps = - arith_simps - add_special - add_0_left - add_0_right - mult_zero_left - mult_zero_right - mult_1_left - mult_1_right - (* simprocs adapted from HOL/ex/Binary.thy *) ML {* local @@ -1694,7 +1692,7 @@ val less = @{term "op < :: int \ int \ bool"} val le = @{term "op \ :: int \ int \ bool"} val simps = @{thms arith_simps} @ @{thms rel_simps} @ - map (fn th => th RS sym) [@{thm numeral_0_eq_0}, @{thm numeral_1_eq_1}] + map (fn th => th RS sym) [@{thm numeral_1_eq_1}] fun prove ctxt goal = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop goal) (K (ALLGOALS (full_simp_tac (HOL_basic_ss addsimps simps)))); fun binary_proc proc ss ct = @@ -1717,14 +1715,25 @@ end *} -simproc_setup binary_int_div ("number_of m div number_of n :: int") = +simproc_setup binary_int_div + ("numeral m div numeral n :: int" | + "numeral m div neg_numeral n :: int" | + "neg_numeral m div numeral n :: int" | + "neg_numeral m div neg_numeral n :: int") = {* K (divmod_proc @{thm int_div_pos_eq} @{thm int_div_neg_eq}) *} -simproc_setup binary_int_mod ("number_of m mod number_of n :: int") = +simproc_setup binary_int_mod + ("numeral m mod numeral n :: int" | + "numeral m mod neg_numeral n :: int" | + "neg_numeral m mod numeral n :: int" | + "neg_numeral m mod neg_numeral n :: int") = {* K (divmod_proc @{thm int_mod_pos_eq} @{thm int_mod_neg_eq}) *} -lemmas posDivAlg_eqn_number_of [simp] = posDivAlg_eqn [of "number_of v" "number_of w"] for v w -lemmas negDivAlg_eqn_number_of [simp] = negDivAlg_eqn [of "number_of v" "number_of w"] for v w +lemmas posDivAlg_eqn_numeral [simp] = + posDivAlg_eqn [of "numeral v" "numeral w", OF zero_less_numeral] for v w + +lemmas negDivAlg_eqn_numeral [simp] = + negDivAlg_eqn [of "numeral v" "neg_numeral w", OF zero_less_numeral] for v w text{*Special-case simplification *} @@ -1741,12 +1750,25 @@ (** The last remaining special cases for constant arithmetic: 1 div z and 1 mod z **) -lemmas div_pos_pos_1_number_of [simp] = div_pos_pos [OF zero_less_one, of "number_of w"] for w -lemmas div_pos_neg_1_number_of [simp] = div_pos_neg [OF zero_less_one, of "number_of w"] for w -lemmas mod_pos_pos_1_number_of [simp] = mod_pos_pos [OF zero_less_one, of "number_of w"] for w -lemmas mod_pos_neg_1_number_of [simp] = mod_pos_neg [OF zero_less_one, of "number_of w"] for w -lemmas posDivAlg_eqn_1_number_of [simp] = posDivAlg_eqn [of concl: 1 "number_of w"] for w -lemmas negDivAlg_eqn_1_number_of [simp] = negDivAlg_eqn [of concl: 1 "number_of w"] for w +lemmas div_pos_pos_1_numeral [simp] = + div_pos_pos [OF zero_less_one, of "numeral w", OF zero_le_numeral] for w + +lemmas div_pos_neg_1_numeral [simp] = + div_pos_neg [OF zero_less_one, of "neg_numeral w", + OF neg_numeral_less_zero] for w + +lemmas mod_pos_pos_1_numeral [simp] = + mod_pos_pos [OF zero_less_one, of "numeral w", OF zero_le_numeral] for w + +lemmas mod_pos_neg_1_numeral [simp] = + mod_pos_neg [OF zero_less_one, of "neg_numeral w", + OF neg_numeral_less_zero] for w + +lemmas posDivAlg_eqn_1_numeral [simp] = + posDivAlg_eqn [of concl: 1 "numeral w", OF zero_less_numeral] for w + +lemmas negDivAlg_eqn_1_numeral [simp] = + negDivAlg_eqn [of concl: 1 "numeral w", OF zero_less_numeral] for w subsubsection {* Monotonicity in the First Argument (Dividend) *} @@ -1928,6 +1950,11 @@ (* REVISIT: should this be generalized to all semiring_div types? *) lemmas zmod_eq_0D [dest!] = zmod_eq_0_iff [THEN iffD1] +lemma zmod_zdiv_equality': + "(m\int) mod n = m - (m div n) * n" + by (rule_tac P="%x. m mod n = x - (m div n) * n" in subst [OF mod_div_equality [of _ n]]) + arith + subsubsection {* Proving @{term "a div (b*c) = (a div b) div c"} *} @@ -1989,6 +2016,26 @@ apply (force simp add: divmod_int_rel_div_mod [THEN zmult2_lemma, THEN divmod_int_rel_mod]) done +lemma div_pos_geq: + fixes k l :: int + assumes "0 < l" and "l \ k" + shows "k div l = (k - l) div l + 1" +proof - + have "k = (k - l) + l" by simp + then obtain j where k: "k = j + l" .. + with assms show ?thesis by simp +qed + +lemma mod_pos_geq: + fixes k l :: int + assumes "0 < l" and "l \ k" + shows "k mod l = (k - l) mod l" +proof - + have "k = (k - l) + l" by simp + then obtain j where k: "k = j + l" .. + with assms show ?thesis by simp +qed + subsubsection {* Splitting Rules for div and mod *} @@ -2046,9 +2093,9 @@ text {* Enable (lin)arith to deal with @{const div} and @{const mod} when these are applied to some constant that is of the form - @{term "number_of k"}: *} -declare split_zdiv [of _ _ "number_of k", arith_split] for k -declare split_zmod [of _ _ "number_of k", arith_split] for k + @{term "numeral k"}: *} +declare split_zdiv [of _ _ "numeral k", arith_split] for k +declare split_zmod [of _ _ "numeral k", arith_split] for k subsubsection {* Speeding up the Division Algorithm with Shifting *} @@ -2090,19 +2137,19 @@ minus_add_distrib [symmetric] mult_minus_right) qed -lemma zdiv_number_of_Bit0 [simp]: - "number_of (Int.Bit0 v) div number_of (Int.Bit0 w) = - number_of v div (number_of w :: int)" -by (simp only: number_of_eq numeral_simps) (simp add: mult_2 [symmetric]) - -lemma zdiv_number_of_Bit1 [simp]: - "number_of (Int.Bit1 v) div number_of (Int.Bit0 w) = - (if (0::int) \ number_of w - then number_of v div (number_of w) - else (number_of v + (1::int)) div (number_of w))" -apply (simp only: number_of_eq numeral_simps UNIV_I split: split_if) -apply (simp add: pos_zdiv_mult_2 neg_zdiv_mult_2 add_ac mult_2 [symmetric]) -done +(* FIXME: add rules for negative numerals *) +lemma zdiv_numeral_Bit0 [simp]: + "numeral (Num.Bit0 v) div numeral (Num.Bit0 w) = + numeral v div (numeral w :: int)" + unfolding numeral.simps unfolding mult_2 [symmetric] + by (rule div_mult_mult1, simp) + +lemma zdiv_numeral_Bit1 [simp]: + "numeral (Num.Bit1 v) div numeral (Num.Bit0 w) = + (numeral v div (numeral w :: int))" + unfolding numeral.simps + unfolding mult_2 [symmetric] add_commute [of _ 1] + by (rule pos_zdiv_mult_2, simp) subsubsection {* Computing mod by Shifting (proofs resemble those for div) *} @@ -2138,24 +2185,19 @@ (simp add: diff_minus add_ac) qed -lemma zmod_number_of_Bit0 [simp]: - "number_of (Int.Bit0 v) mod number_of (Int.Bit0 w) = - (2::int) * (number_of v mod number_of w)" -apply (simp only: number_of_eq numeral_simps) -apply (simp add: mod_mult_mult1 pos_zmod_mult_2 - neg_zmod_mult_2 add_ac mult_2 [symmetric]) -done - -lemma zmod_number_of_Bit1 [simp]: - "number_of (Int.Bit1 v) mod number_of (Int.Bit0 w) = - (if (0::int) \ number_of w - then 2 * (number_of v mod number_of w) + 1 - else 2 * ((number_of v + (1::int)) mod number_of w) - 1)" -apply (simp only: number_of_eq numeral_simps) -apply (simp add: mod_mult_mult1 pos_zmod_mult_2 - neg_zmod_mult_2 add_ac mult_2 [symmetric]) -done - +(* FIXME: add rules for negative numerals *) +lemma zmod_numeral_Bit0 [simp]: + "numeral (Num.Bit0 v) mod numeral (Num.Bit0 w) = + (2::int) * (numeral v mod numeral w)" + unfolding numeral_Bit0 [of v] numeral_Bit0 [of w] + unfolding mult_2 [symmetric] by (rule mod_mult_mult1) + +lemma zmod_numeral_Bit1 [simp]: + "numeral (Num.Bit1 v) mod numeral (Num.Bit0 w) = + 2 * (numeral v mod numeral w) + (1::int)" + unfolding numeral_Bit1 [of v] numeral_Bit0 [of w] + unfolding mult_2 [symmetric] add_commute [of _ 1] + by (rule pos_zmod_mult_2, simp) lemma zdiv_eq_0_iff: "(i::int) div k = 0 \ k=0 \ 0\i \ i i\0 \ k f dvd (n::int) ==> f dvd m mod n" by (rule dvd_mod) (* TODO: remove *) @@ -2242,6 +2287,12 @@ lemma zdvd_zmod_imp_zdvd: "k dvd m mod n ==> k dvd n ==> k dvd (m::int)" by (rule dvd_mod_imp_dvd) (* TODO: remove *) +lemmas dvd_eq_mod_eq_0_numeral [simp] = + dvd_eq_mod_eq_0 [of "numeral x" "numeral y"] for x y + + +subsubsection {* Further properties *} + lemma zmult_div_cancel: "(n::int) * (m div n) = m - (m mod n)" using zmod_zdiv_equality[where a="m" and b="n"] by (simp add: algebra_simps) @@ -2408,42 +2459,31 @@ thus ?lhs by simp qed -lemma div_nat_number_of [simp]: - "(number_of v :: nat) div number_of v' = - (if neg (number_of v :: int) then 0 - else nat (number_of v div number_of v'))" - unfolding nat_number_of_def number_of_is_id neg_def +lemma div_nat_numeral [simp]: + "(numeral v :: nat) div numeral v' = nat (numeral v div numeral v')" by (simp add: nat_div_distrib) -lemma one_div_nat_number_of [simp]: - "Suc 0 div number_of v' = nat (1 div number_of v')" - by (simp del: semiring_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric] semiring_numeral_1_eq_1 [symmetric]) - -lemma mod_nat_number_of [simp]: - "(number_of v :: nat) mod number_of v' = - (if neg (number_of v :: int) then 0 - else if neg (number_of v' :: int) then number_of v - else nat (number_of v mod number_of v'))" - unfolding nat_number_of_def number_of_is_id neg_def +lemma one_div_nat_numeral [simp]: + "Suc 0 div numeral v' = nat (1 div numeral v')" + by (subst nat_div_distrib, simp_all) + +lemma mod_nat_numeral [simp]: + "(numeral v :: nat) mod numeral v' = nat (numeral v mod numeral v')" by (simp add: nat_mod_distrib) -lemma one_mod_nat_number_of [simp]: - "Suc 0 mod number_of v' = - (if neg (number_of v' :: int) then Suc 0 - else nat (1 mod number_of v'))" -by (simp del: semiring_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric] semiring_numeral_1_eq_1 [symmetric]) - -lemmas dvd_eq_mod_eq_0_number_of [simp] = - dvd_eq_mod_eq_0 [of "number_of x" "number_of y"] for x y - - -subsubsection {* Nitpick *} - -lemma zmod_zdiv_equality': -"(m\int) mod n = m - (m div n) * n" -by (rule_tac P="%x. m mod n = x - (m div n) * n" - in subst [OF mod_div_equality [of _ n]]) - arith +lemma one_mod_nat_numeral [simp]: + "Suc 0 mod numeral v' = nat (1 mod numeral v')" + by (subst nat_mod_distrib) simp_all + +lemma mod_2_not_eq_zero_eq_one_int: + fixes k :: int + shows "k mod 2 \ 0 \ k mod 2 = 1" + by auto + + +subsubsection {* Tools setup *} + +text {* Nitpick *} lemmas [nitpick_unfold] = dvd_eq_mod_eq_0 mod_div_equality' zmod_zdiv_equality' @@ -2461,7 +2501,7 @@ apsnd ((op *) (sgn l)) (if 0 < l \ 0 \ k \ l < 0 \ k < 0 then pdivmod k l else (let (r, s) = pdivmod k l in - if s = 0 then (- r, 0) else (- r - 1, \l\ - s))))" + if s = 0 then (- r, 0) else (- r - 1, \l\ - s))))" proof - have aux: "\q::int. - k = l * q \ k = l * - q" by auto show ?thesis @@ -2481,45 +2521,6 @@ then show ?thesis by (simp add: divmod_int_pdivmod) qed -context ring_1 -begin - -lemma of_int_num [code]: - "of_int k = (if k = 0 then 0 else if k < 0 then - - of_int (- k) else let - (l, m) = divmod_int k 2; - l' = of_int l - in if m = 0 then l' + l' else l' + l' + 1)" -proof - - have aux1: "k mod (2\int) \ (0\int) \ - of_int k = of_int (k div 2 * 2 + 1)" - proof - - have "k mod 2 < 2" by (auto intro: pos_mod_bound) - moreover have "0 \ k mod 2" by (auto intro: pos_mod_sign) - moreover assume "k mod 2 \ 0" - ultimately have "k mod 2 = 1" by arith - moreover have "of_int k = of_int (k div 2 * 2 + k mod 2)" by simp - ultimately show ?thesis by auto - qed - have aux2: "\x. of_int 2 * x = x + x" - proof - - fix x - have int2: "(2::int) = 1 + 1" by arith - show "of_int 2 * x = x + x" - unfolding int2 of_int_add left_distrib by simp - qed - have aux3: "\x. x * of_int 2 = x + x" - proof - - fix x - have int2: "(2::int) = 1 + 1" by arith - show "x * of_int 2 = x + x" - unfolding int2 of_int_add right_distrib by simp - qed - from aux1 show ?thesis by (auto simp add: divmod_int_mod_div Let_def aux2 aux3) -qed - -end - code_modulename SML Divides Arith diff -r 29e92b644d6c -r f067afe98049 src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy --- a/src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy Mon Mar 26 10:56:56 2012 +0200 @@ -6,7 +6,7 @@ theory Imperative_Quicksort imports - Imperative_HOL + "~~/src/HOL/Imperative_HOL/Imperative_HOL" Subarray "~~/src/HOL/Library/Multiset" "~~/src/HOL/Library/Efficient_Nat" @@ -593,8 +593,8 @@ proof (induct a l r p arbitrary: h rule: part1.induct) case (1 a l r p) thus ?case unfolding part1.simps [of a l r] - apply (auto intro!: success_intros del: success_ifI simp add: not_le) - apply (auto intro!: effect_intros effect_swapI) + apply (auto intro!: success_intros simp add: not_le) + apply (auto intro!: effect_intros) done qed diff -r 29e92b644d6c -r f067afe98049 src/HOL/Imperative_HOL/ex/Imperative_Reverse.thy --- a/src/HOL/Imperative_HOL/ex/Imperative_Reverse.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Imperative_HOL/ex/Imperative_Reverse.thy Mon Mar 26 10:56:56 2012 +0200 @@ -5,7 +5,7 @@ header {* An imperative in-place reversal on arrays *} theory Imperative_Reverse -imports Subarray Imperative_HOL +imports Subarray "~~/src/HOL/Imperative_HOL/Imperative_HOL" begin fun swap :: "'a\heap array \ nat \ nat \ unit Heap" where @@ -107,7 +107,7 @@ shows "Array.get h' a = List.rev (Array.get h a)" using rev2_rev'[OF assms] rev_length[OF assms] assms by (cases "Array.length h a = 0", auto simp add: Array.length_def - subarray_def sublist'_all rev.simps[where j=0] elim!: effect_elims) + subarray_def rev.simps[where j=0] elim!: effect_elims) (drule sym[of "List.length (Array.get h a)"], simp) definition "example = (Array.make 10 id \= (\a. rev a 0 9))" @@ -115,3 +115,4 @@ export_code example checking SML SML_imp OCaml? OCaml_imp? Haskell? Scala? Scala_imp? end + diff -r 29e92b644d6c -r f067afe98049 src/HOL/Imperative_HOL/ex/SatChecker.thy --- a/src/HOL/Imperative_HOL/ex/SatChecker.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Imperative_HOL/ex/SatChecker.thy Mon Mar 26 10:56:56 2012 +0200 @@ -702,15 +702,7 @@ else raise(''No empty clause'')) }" -section {* Code generation setup *} - -code_type ProofStep - (SML "MinisatProofStep.ProofStep") - -code_const ProofDone and Root and Conflict and Delete and Xstep - (SML "MinisatProofStep.ProofDone" and "MinisatProofStep.Root ((_),/ (_))" and "MinisatProofStep.Conflict ((_),/ (_))" and "MinisatProofStep.Delete" and "MinisatProofStep.Xstep ((_),/ (_))") - -export_code checker tchecker lchecker in SML +export_code checker tchecker lchecker checking SML end diff -r 29e92b644d6c -r f067afe98049 src/HOL/Imperative_HOL/ex/Subarray.thy --- a/src/HOL/Imperative_HOL/ex/Subarray.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Imperative_HOL/ex/Subarray.thy Mon Mar 26 10:56:56 2012 +0200 @@ -5,7 +5,7 @@ header {* Theorems about sub arrays *} theory Subarray -imports Array Sublist +imports "~~/src/HOL/Imperative_HOL/Array" Sublist begin definition subarray :: "nat \ nat \ ('a::heap) array \ heap \ 'a list" where diff -r 29e92b644d6c -r f067afe98049 src/HOL/Import/HOL_Light/HOLLightInt.thy --- a/src/HOL/Import/HOL_Light/HOLLightInt.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Import/HOL_Light/HOLLightInt.thy Mon Mar 26 10:56:56 2012 +0200 @@ -40,7 +40,7 @@ lemma DEF_int_mul: "op * = (\u ua. floor (real u * real ua))" - by (metis floor_number_of number_of_is_id number_of_real_def real_eq_of_int real_of_int_mult) + by (metis floor_real_of_int real_of_int_mult) lemma DEF_int_abs: "abs = (\u. floor (abs (real u)))" @@ -72,7 +72,7 @@ lemma INT_IMAGE: "(\n. x = int n) \ (\n. x = - int n)" - by (metis number_of_eq number_of_is_id of_int_of_nat) + by (metis of_int_eq_id id_def of_int_of_nat) lemma DEF_int_pow: "op ^ = (\u ua. floor (real u ^ ua))" diff -r 29e92b644d6c -r f067afe98049 src/HOL/Int.thy --- a/src/HOL/Int.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Int.thy Mon Mar 26 10:56:56 2012 +0200 @@ -6,10 +6,9 @@ header {* The Integers as Equivalence Classes over Pairs of Natural Numbers *} theory Int -imports Equiv_Relations Nat Wellfounded +imports Equiv_Relations Wellfounded uses ("Tools/numeral.ML") - ("Tools/numeral_syntax.ML") ("Tools/int_arith.ML") begin @@ -323,15 +322,20 @@ lemma of_int_of_nat_eq [simp]: "of_int (int n) = of_nat n" by (induct n) auto +lemma of_int_numeral [simp, code_post]: "of_int (numeral k) = numeral k" + by (simp add: of_nat_numeral [symmetric] of_int_of_nat_eq [symmetric]) + +lemma of_int_neg_numeral [simp, code_post]: "of_int (neg_numeral k) = neg_numeral k" + unfolding neg_numeral_def neg_numeral_class.neg_numeral_def + by (simp only: of_int_minus of_int_numeral) + lemma of_int_power: "of_int (z ^ n) = of_int z ^ n" by (induct n) simp_all end -text{*Class for unital rings with characteristic zero. - Includes non-ordered rings like the complex numbers.*} -class ring_char_0 = ring_1 + semiring_char_0 +context ring_char_0 begin lemma of_int_eq_iff [simp]: @@ -579,230 +583,27 @@ apply (simp add: int_def minus add diff_minus) done - -subsection {* Binary representation *} - -text {* - This formalization defines binary arithmetic in terms of the integers - rather than using a datatype. This avoids multiple representations (leading - zeroes, etc.) See @{text "ZF/Tools/twos-compl.ML"}, function @{text - int_of_binary}, for the numerical interpretation. - - The representation expects that @{text "(m mod 2)"} is 0 or 1, - even if m is negative; - For instance, @{text "-5 div 2 = -3"} and @{text "-5 mod 2 = 1"}; thus - @{text "-5 = (-3)*2 + 1"}. - - This two's complement binary representation derives from the paper - "An Efficient Representation of Arithmetic for Term Rewriting" by - Dave Cohen and Phil Watson, Rewriting Techniques and Applications, - Springer LNCS 488 (240-251), 1991. -*} - -subsubsection {* The constructors @{term Bit0}, @{term Bit1}, @{term Pls} and @{term Min} *} - -definition Pls :: int where - "Pls = 0" +lemma Let_numeral [simp]: "Let (numeral v) f = f (numeral v)" + -- {* Unfold all @{text let}s involving constants *} + unfolding Let_def .. -definition Min :: int where - "Min = - 1" - -definition Bit0 :: "int \ int" where - "Bit0 k = k + k" - -definition Bit1 :: "int \ int" where - "Bit1 k = 1 + k + k" - -class number = -- {* for numeric types: nat, int, real, \dots *} - fixes number_of :: "int \ 'a" - -use "Tools/numeral.ML" - -syntax - "_Numeral" :: "num_const \ 'a" ("_") - -use "Tools/numeral_syntax.ML" -setup Numeral_Syntax.setup - -abbreviation - "Numeral0 \ number_of Pls" - -abbreviation - "Numeral1 \ number_of (Bit1 Pls)" - -lemma Let_number_of [simp]: "Let (number_of v) f = f (number_of v)" +lemma Let_neg_numeral [simp]: "Let (neg_numeral v) f = f (neg_numeral v)" -- {* Unfold all @{text let}s involving constants *} unfolding Let_def .. -definition succ :: "int \ int" where - "succ k = k + 1" - -definition pred :: "int \ int" where - "pred k = k - 1" - -lemmas max_number_of [simp] = max_def [of "number_of u" "number_of v"] - and min_number_of [simp] = min_def [of "number_of u" "number_of v"] - for u v - -- {* unfolding @{text minx} and @{text max} on numerals *} - -lemmas numeral_simps = - succ_def pred_def Pls_def Min_def Bit0_def Bit1_def - -text {* Removal of leading zeroes *} - -lemma Bit0_Pls [simp, code_post]: - "Bit0 Pls = Pls" - unfolding numeral_simps by simp - -lemma Bit1_Min [simp, code_post]: - "Bit1 Min = Min" - unfolding numeral_simps by simp - -lemmas normalize_bin_simps = - Bit0_Pls Bit1_Min - - -subsubsection {* Successor and predecessor functions *} - -text {* Successor *} - -lemma succ_Pls: - "succ Pls = Bit1 Pls" - unfolding numeral_simps by simp - -lemma succ_Min: - "succ Min = Pls" - unfolding numeral_simps by simp - -lemma succ_Bit0: - "succ (Bit0 k) = Bit1 k" - unfolding numeral_simps by simp - -lemma succ_Bit1: - "succ (Bit1 k) = Bit0 (succ k)" - unfolding numeral_simps by simp - -lemmas succ_bin_simps [simp] = - succ_Pls succ_Min succ_Bit0 succ_Bit1 - -text {* Predecessor *} - -lemma pred_Pls: - "pred Pls = Min" - unfolding numeral_simps by simp - -lemma pred_Min: - "pred Min = Bit0 Min" - unfolding numeral_simps by simp - -lemma pred_Bit0: - "pred (Bit0 k) = Bit1 (pred k)" - unfolding numeral_simps by simp - -lemma pred_Bit1: - "pred (Bit1 k) = Bit0 k" - unfolding numeral_simps by simp - -lemmas pred_bin_simps [simp] = - pred_Pls pred_Min pred_Bit0 pred_Bit1 - - -subsubsection {* Binary arithmetic *} - -text {* Addition *} - -lemma add_Pls: - "Pls + k = k" - unfolding numeral_simps by simp - -lemma add_Min: - "Min + k = pred k" - unfolding numeral_simps by simp +text {* Unfold @{text min} and @{text max} on numerals. *} -lemma add_Bit0_Bit0: - "(Bit0 k) + (Bit0 l) = Bit0 (k + l)" - unfolding numeral_simps by simp - -lemma add_Bit0_Bit1: - "(Bit0 k) + (Bit1 l) = Bit1 (k + l)" - unfolding numeral_simps by simp - -lemma add_Bit1_Bit0: - "(Bit1 k) + (Bit0 l) = Bit1 (k + l)" - unfolding numeral_simps by simp - -lemma add_Bit1_Bit1: - "(Bit1 k) + (Bit1 l) = Bit0 (k + succ l)" - unfolding numeral_simps by simp - -lemma add_Pls_right: - "k + Pls = k" - unfolding numeral_simps by simp - -lemma add_Min_right: - "k + Min = pred k" - unfolding numeral_simps by simp - -lemmas add_bin_simps [simp] = - add_Pls add_Min add_Pls_right add_Min_right - add_Bit0_Bit0 add_Bit0_Bit1 add_Bit1_Bit0 add_Bit1_Bit1 - -text {* Negation *} - -lemma minus_Pls: - "- Pls = Pls" - unfolding numeral_simps by simp - -lemma minus_Min: - "- Min = Bit1 Pls" - unfolding numeral_simps by simp - -lemma minus_Bit0: - "- (Bit0 k) = Bit0 (- k)" - unfolding numeral_simps by simp +lemmas max_number_of [simp] = + max_def [of "numeral u" "numeral v"] + max_def [of "numeral u" "neg_numeral v"] + max_def [of "neg_numeral u" "numeral v"] + max_def [of "neg_numeral u" "neg_numeral v"] for u v -lemma minus_Bit1: - "- (Bit1 k) = Bit1 (pred (- k))" - unfolding numeral_simps by simp - -lemmas minus_bin_simps [simp] = - minus_Pls minus_Min minus_Bit0 minus_Bit1 - -text {* Subtraction *} - -lemma diff_bin_simps [simp]: - "k - Pls = k" - "k - Min = succ k" - "Pls - (Bit0 l) = Bit0 (Pls - l)" - "Pls - (Bit1 l) = Bit1 (Min - l)" - "Min - (Bit0 l) = Bit1 (Min - l)" - "Min - (Bit1 l) = Bit0 (Min - l)" - "(Bit0 k) - (Bit0 l) = Bit0 (k - l)" - "(Bit0 k) - (Bit1 l) = Bit1 (pred k - l)" - "(Bit1 k) - (Bit0 l) = Bit1 (k - l)" - "(Bit1 k) - (Bit1 l) = Bit0 (k - l)" - unfolding numeral_simps by simp_all - -text {* Multiplication *} - -lemma mult_Pls: - "Pls * w = Pls" - unfolding numeral_simps by simp - -lemma mult_Min: - "Min * k = - k" - unfolding numeral_simps by simp - -lemma mult_Bit0: - "(Bit0 k) * l = Bit0 (k * l)" - unfolding numeral_simps int_distrib by simp - -lemma mult_Bit1: - "(Bit1 k) * l = (Bit0 (k * l)) + l" - unfolding numeral_simps int_distrib by simp - -lemmas mult_bin_simps [simp] = - mult_Pls mult_Min mult_Bit0 mult_Bit1 +lemmas min_number_of [simp] = + min_def [of "numeral u" "numeral v"] + min_def [of "numeral u" "neg_numeral v"] + min_def [of "neg_numeral u" "numeral v"] + min_def [of "neg_numeral u" "neg_numeral v"] for u v subsubsection {* Binary comparisons *} @@ -812,7 +613,7 @@ lemma even_less_0_iff: "a + a < 0 \ a < (0::'a::linordered_idom)" proof - - have "a + a < 0 \ (1+1)*a < 0" by (simp add: left_distrib) + have "a + a < 0 \ (1+1)*a < 0" by (simp add: left_distrib del: one_add_one) also have "(1+1)*a < 0 \ a < 0" by (simp add: mult_less_0_iff zero_less_two order_less_not_sym [OF zero_less_two]) @@ -824,7 +625,7 @@ shows "(0::int) < 1 + z" proof - have "0 \ z" by fact - also have "... < z + 1" by (rule less_add_one) + also have "... < z + 1" by (rule less_add_one) also have "... = 1 + z" by (simp add: add_ac) finally show "0 < 1 + z" . qed @@ -841,276 +642,6 @@ add: algebra_simps of_nat_1 [where 'a=int, symmetric] of_nat_add [symmetric]) qed -lemma bin_less_0_simps: - "Pls < 0 \ False" - "Min < 0 \ True" - "Bit0 w < 0 \ w < 0" - "Bit1 w < 0 \ w < 0" - unfolding numeral_simps - by (simp_all add: even_less_0_iff odd_less_0_iff) - -lemma less_bin_lemma: "k < l \ k - l < (0::int)" - by simp - -lemma le_iff_pred_less: "k \ l \ pred k < l" - unfolding numeral_simps - proof - have "k - 1 < k" by simp - also assume "k \ l" - finally show "k - 1 < l" . - next - assume "k - 1 < l" - hence "(k - 1) + 1 \ l" by (rule zless_imp_add1_zle) - thus "k \ l" by simp - qed - -lemma succ_pred: "succ (pred x) = x" - unfolding numeral_simps by simp - -text {* Less-than *} - -lemma less_bin_simps [simp]: - "Pls < Pls \ False" - "Pls < Min \ False" - "Pls < Bit0 k \ Pls < k" - "Pls < Bit1 k \ Pls \ k" - "Min < Pls \ True" - "Min < Min \ False" - "Min < Bit0 k \ Min < k" - "Min < Bit1 k \ Min < k" - "Bit0 k < Pls \ k < Pls" - "Bit0 k < Min \ k \ Min" - "Bit1 k < Pls \ k < Pls" - "Bit1 k < Min \ k < Min" - "Bit0 k < Bit0 l \ k < l" - "Bit0 k < Bit1 l \ k \ l" - "Bit1 k < Bit0 l \ k < l" - "Bit1 k < Bit1 l \ k < l" - unfolding le_iff_pred_less - less_bin_lemma [of Pls] - less_bin_lemma [of Min] - less_bin_lemma [of "k"] - less_bin_lemma [of "Bit0 k"] - less_bin_lemma [of "Bit1 k"] - less_bin_lemma [of "pred Pls"] - less_bin_lemma [of "pred k"] - by (simp_all add: bin_less_0_simps succ_pred) - -text {* Less-than-or-equal *} - -lemma le_bin_simps [simp]: - "Pls \ Pls \ True" - "Pls \ Min \ False" - "Pls \ Bit0 k \ Pls \ k" - "Pls \ Bit1 k \ Pls \ k" - "Min \ Pls \ True" - "Min \ Min \ True" - "Min \ Bit0 k \ Min < k" - "Min \ Bit1 k \ Min \ k" - "Bit0 k \ Pls \ k \ Pls" - "Bit0 k \ Min \ k \ Min" - "Bit1 k \ Pls \ k < Pls" - "Bit1 k \ Min \ k \ Min" - "Bit0 k \ Bit0 l \ k \ l" - "Bit0 k \ Bit1 l \ k \ l" - "Bit1 k \ Bit0 l \ k < l" - "Bit1 k \ Bit1 l \ k \ l" - unfolding not_less [symmetric] - by (simp_all add: not_le) - -text {* Equality *} - -lemma eq_bin_simps [simp]: - "Pls = Pls \ True" - "Pls = Min \ False" - "Pls = Bit0 l \ Pls = l" - "Pls = Bit1 l \ False" - "Min = Pls \ False" - "Min = Min \ True" - "Min = Bit0 l \ False" - "Min = Bit1 l \ Min = l" - "Bit0 k = Pls \ k = Pls" - "Bit0 k = Min \ False" - "Bit1 k = Pls \ False" - "Bit1 k = Min \ k = Min" - "Bit0 k = Bit0 l \ k = l" - "Bit0 k = Bit1 l \ False" - "Bit1 k = Bit0 l \ False" - "Bit1 k = Bit1 l \ k = l" - unfolding order_eq_iff [where 'a=int] - by (simp_all add: not_less) - - -subsection {* Converting Numerals to Rings: @{term number_of} *} - -class number_ring = number + comm_ring_1 + - assumes number_of_eq: "number_of k = of_int k" - -class number_semiring = number + comm_semiring_1 + - assumes number_of_int: "number_of (int n) = of_nat n" - -instance number_ring \ number_semiring -proof - fix n show "number_of (int n) = (of_nat n :: 'a)" - unfolding number_of_eq by (rule of_int_of_nat_eq) -qed - -text {* self-embedding of the integers *} - -instantiation int :: number_ring -begin - -definition - int_number_of_def: "number_of w = (of_int w \ int)" - -instance proof -qed (simp only: int_number_of_def) - -end - -lemma number_of_is_id: - "number_of (k::int) = k" - unfolding int_number_of_def by simp - -lemma number_of_succ: - "number_of (succ k) = (1 + number_of k ::'a::number_ring)" - unfolding number_of_eq numeral_simps by simp - -lemma number_of_pred: - "number_of (pred w) = (- 1 + number_of w ::'a::number_ring)" - unfolding number_of_eq numeral_simps by simp - -lemma number_of_minus: - "number_of (uminus w) = (- (number_of w)::'a::number_ring)" - unfolding number_of_eq by (rule of_int_minus) - -lemma number_of_add: - "number_of (v + w) = (number_of v + number_of w::'a::number_ring)" - unfolding number_of_eq by (rule of_int_add) - -lemma number_of_diff: - "number_of (v - w) = (number_of v - number_of w::'a::number_ring)" - unfolding number_of_eq by (rule of_int_diff) - -lemma number_of_mult: - "number_of (v * w) = (number_of v * number_of w::'a::number_ring)" - unfolding number_of_eq by (rule of_int_mult) - -text {* - The correctness of shifting. - But it doesn't seem to give a measurable speed-up. -*} - -lemma double_number_of_Bit0: - "(1 + 1) * number_of w = (number_of (Bit0 w) ::'a::number_ring)" - unfolding number_of_eq numeral_simps left_distrib by simp - -text {* - Converting numerals 0 and 1 to their abstract versions. -*} - -lemma semiring_numeral_0_eq_0 [simp, code_post]: - "Numeral0 = (0::'a::number_semiring)" - using number_of_int [where 'a='a and n=0] - unfolding numeral_simps by simp - -lemma semiring_numeral_1_eq_1 [simp, code_post]: - "Numeral1 = (1::'a::number_semiring)" - using number_of_int [where 'a='a and n=1] - unfolding numeral_simps by simp - -lemma numeral_0_eq_0: (* FIXME delete candidate *) - "Numeral0 = (0::'a::number_ring)" - by (rule semiring_numeral_0_eq_0) - -lemma numeral_1_eq_1: (* FIXME delete candidate *) - "Numeral1 = (1::'a::number_ring)" - by (rule semiring_numeral_1_eq_1) - -text {* - Special-case simplification for small constants. -*} - -text{* - Unary minus for the abstract constant 1. Cannot be inserted - as a simprule until later: it is @{text number_of_Min} re-oriented! -*} - -lemma numeral_m1_eq_minus_1: - "(-1::'a::number_ring) = - 1" - unfolding number_of_eq numeral_simps by simp - -lemma mult_minus1 [simp]: - "-1 * z = -(z::'a::number_ring)" - unfolding number_of_eq numeral_simps by simp - -lemma mult_minus1_right [simp]: - "z * -1 = -(z::'a::number_ring)" - unfolding number_of_eq numeral_simps by simp - -(*Negation of a coefficient*) -lemma minus_number_of_mult [simp]: - "- (number_of w) * z = number_of (uminus w) * (z::'a::number_ring)" - unfolding number_of_eq by simp - -text {* Subtraction *} - -lemma diff_number_of_eq: - "number_of v - number_of w = - (number_of (v + uminus w)::'a::number_ring)" - unfolding number_of_eq by simp - -lemma number_of_Pls: - "number_of Pls = (0::'a::number_ring)" - unfolding number_of_eq numeral_simps by simp - -lemma number_of_Min: - "number_of Min = (- 1::'a::number_ring)" - unfolding number_of_eq numeral_simps by simp - -lemma number_of_Bit0: - "number_of (Bit0 w) = (0::'a::number_ring) + (number_of w) + (number_of w)" - unfolding number_of_eq numeral_simps by simp - -lemma number_of_Bit1: - "number_of (Bit1 w) = (1::'a::number_ring) + (number_of w) + (number_of w)" - unfolding number_of_eq numeral_simps by simp - - -subsubsection {* Equality of Binary Numbers *} - -text {* First version by Norbert Voelker *} - -definition (*for simplifying equalities*) iszero :: "'a\semiring_1 \ bool" where - "iszero z \ z = 0" - -lemma iszero_0: "iszero 0" - by (simp add: iszero_def) - -lemma iszero_Numeral0: "iszero (Numeral0 :: 'a::number_ring)" - by (simp add: iszero_0) - -lemma not_iszero_1: "\ iszero 1" - by (simp add: iszero_def) - -lemma not_iszero_Numeral1: "\ iszero (Numeral1 :: 'a::number_ring)" - by (simp add: not_iszero_1) - -lemma eq_number_of_eq [simp]: - "((number_of x::'a::number_ring) = number_of y) = - iszero (number_of (x + uminus y) :: 'a)" -unfolding iszero_def number_of_add number_of_minus -by (simp add: algebra_simps) - -lemma iszero_number_of_Pls: - "iszero ((number_of Pls)::'a::number_ring)" -unfolding iszero_def numeral_0_eq_0 .. - -lemma nonzero_number_of_Min: - "~ iszero ((number_of Min)::'a::number_ring)" -unfolding iszero_def numeral_m1_eq_minus_1 by simp - - subsubsection {* Comparisons, for Ordered Rings *} lemmas double_eq_0_iff = double_zero @@ -1137,129 +668,6 @@ qed qed -lemma iszero_number_of_Bit0: - "iszero (number_of (Bit0 w)::'a) = - iszero (number_of w::'a::{ring_char_0,number_ring})" -proof - - have "(of_int w + of_int w = (0::'a)) \ (w = 0)" - proof - - assume eq: "of_int w + of_int w = (0::'a)" - then have "of_int (w + w) = (of_int 0 :: 'a)" by simp - then have "w + w = 0" by (simp only: of_int_eq_iff) - then show "w = 0" by (simp only: double_eq_0_iff) - qed - thus ?thesis - by (auto simp add: iszero_def number_of_eq numeral_simps) -qed - -lemma iszero_number_of_Bit1: - "~ iszero (number_of (Bit1 w)::'a::{ring_char_0,number_ring})" -proof - - have "1 + of_int w + of_int w \ (0::'a)" - proof - assume eq: "1 + of_int w + of_int w = (0::'a)" - hence "of_int (1 + w + w) = (of_int 0 :: 'a)" by simp - hence "1 + w + w = 0" by (simp only: of_int_eq_iff) - with odd_nonzero show False by blast - qed - thus ?thesis - by (auto simp add: iszero_def number_of_eq numeral_simps) -qed - -lemmas iszero_simps [simp] = - iszero_0 not_iszero_1 - iszero_number_of_Pls nonzero_number_of_Min - iszero_number_of_Bit0 iszero_number_of_Bit1 -(* iszero_number_of_Pls would never normally be used - because its lhs simplifies to "iszero 0" *) - -text {* Less-Than or Equals *} - -text {* Reduces @{term "a\b"} to @{term "~ (b x < y" - unfolding number_of_eq by (rule of_int_less_iff) - -lemma le_number_of [simp]: - "(number_of x::'a::{linordered_idom,number_ring}) \ number_of y \ x \ y" - unfolding number_of_eq by (rule of_int_le_iff) - -lemma eq_number_of [simp]: - "(number_of x::'a::{ring_char_0,number_ring}) = number_of y \ x = y" - unfolding number_of_eq by (rule of_int_eq_iff) - -lemmas rel_simps = - less_number_of less_bin_simps - le_number_of le_bin_simps - eq_number_of_eq eq_bin_simps - iszero_simps - - -subsubsection {* Simplification of arithmetic when nested to the right. *} - -lemma add_number_of_left [simp]: - "number_of v + (number_of w + z) = - (number_of(v + w) + z::'a::number_ring)" - by (simp add: add_assoc [symmetric]) - -lemma mult_number_of_left [simp]: - "number_of v * (number_of w * z) = - (number_of(v * w) * z::'a::number_ring)" - by (simp add: mult_assoc [symmetric]) - -lemma add_number_of_diff1: - "number_of v + (number_of w - c) = - number_of(v + w) - (c::'a::number_ring)" - by (simp add: diff_minus) - -lemma add_number_of_diff2 [simp]: - "number_of v + (c - number_of w) = - number_of (v + uminus w) + (c::'a::number_ring)" -by (simp add: algebra_simps diff_number_of_eq [symmetric]) - - - subsection {* The Set of Integers *} @@ -1363,14 +771,8 @@ qed qed -lemma Ints_number_of [simp]: - "(number_of w :: 'a::number_ring) \ Ints" - unfolding number_of_eq Ints_def by simp - -lemma Nats_number_of [simp]: - "Int.Pls \ w \ (number_of w :: 'a::number_ring) \ Nats" -unfolding Int.Pls_def number_of_eq -by (simp only: of_nat_nat [symmetric] of_nat_in_Nats) +lemma Nats_numeral [simp]: "numeral w \ Nats" + using of_nat_in_Nats [of "numeral w"] by simp lemma Ints_odd_less_0: assumes in_Ints: "a \ Ints" @@ -1412,100 +814,16 @@ lemmas int_setprod = of_nat_setprod [where 'a=int] -subsection{*Inequality Reasoning for the Arithmetic Simproc*} - -lemma add_numeral_0: "Numeral0 + a = (a::'a::number_ring)" -by simp - -lemma add_numeral_0_right: "a + Numeral0 = (a::'a::number_ring)" -by simp - -lemma mult_numeral_1: "Numeral1 * a = (a::'a::number_ring)" -by simp - -lemma mult_numeral_1_right: "a * Numeral1 = (a::'a::number_ring)" -by simp - -lemma divide_numeral_1: "a / Numeral1 = (a::'a::{number_ring,field})" -by simp - -lemma inverse_numeral_1: - "inverse Numeral1 = (Numeral1::'a::{number_ring,field})" -by simp - -text{*Theorem lists for the cancellation simprocs. The use of binary numerals -for 0 and 1 reduces the number of special cases.*} - -lemmas add_0s = add_numeral_0 add_numeral_0_right -lemmas mult_1s = mult_numeral_1 mult_numeral_1_right - mult_minus1 mult_minus1_right - - -subsection{*Special Arithmetic Rules for Abstract 0 and 1*} - -text{*Arithmetic computations are defined for binary literals, which leaves 0 -and 1 as special cases. Addition already has rules for 0, but not 1. -Multiplication and unary minus already have rules for both 0 and 1.*} - - -lemma binop_eq: "[|f x y = g x y; x = x'; y = y'|] ==> f x' y' = g x' y'" -by simp - - -lemmas add_number_of_eq = number_of_add [symmetric] - -text{*Allow 1 on either or both sides*} -lemma semiring_one_add_one_is_two: "1 + 1 = (2::'a::number_semiring)" - using number_of_int [where 'a='a and n="Suc (Suc 0)"] - by (simp add: numeral_simps) - -lemma one_add_one_is_two: "1 + 1 = (2::'a::number_ring)" -by (rule semiring_one_add_one_is_two) - -lemmas add_special = - one_add_one_is_two - binop_eq [of "op +", OF add_number_of_eq numeral_1_eq_1 refl] - binop_eq [of "op +", OF add_number_of_eq refl numeral_1_eq_1] - -text{*Allow 1 on either or both sides (1-1 already simplifies to 0)*} -lemmas diff_special = - binop_eq [of "op -", OF diff_number_of_eq numeral_1_eq_1 refl] - binop_eq [of "op -", OF diff_number_of_eq refl numeral_1_eq_1] - -text{*Allow 0 or 1 on either side with a binary numeral on the other*} -lemmas eq_special = - binop_eq [of "op =", OF eq_number_of_eq numeral_0_eq_0 refl] - binop_eq [of "op =", OF eq_number_of_eq numeral_1_eq_1 refl] - binop_eq [of "op =", OF eq_number_of_eq refl numeral_0_eq_0] - binop_eq [of "op =", OF eq_number_of_eq refl numeral_1_eq_1] - -text{*Allow 0 or 1 on either side with a binary numeral on the other*} -lemmas less_special = - binop_eq [of "op <", OF less_number_of numeral_0_eq_0 refl] - binop_eq [of "op <", OF less_number_of numeral_1_eq_1 refl] - binop_eq [of "op <", OF less_number_of refl numeral_0_eq_0] - binop_eq [of "op <", OF less_number_of refl numeral_1_eq_1] - -text{*Allow 0 or 1 on either side with a binary numeral on the other*} -lemmas le_special = - binop_eq [of "op \", OF le_number_of numeral_0_eq_0 refl] - binop_eq [of "op \", OF le_number_of numeral_1_eq_1 refl] - binop_eq [of "op \", OF le_number_of refl numeral_0_eq_0] - binop_eq [of "op \", OF le_number_of refl numeral_1_eq_1] - -lemmas arith_special[simp] = - add_special diff_special eq_special less_special le_special - - text {* Legacy theorems *} lemmas zle_int = of_nat_le_iff [where 'a=int] lemmas int_int_eq = of_nat_eq_iff [where 'a=int] +lemmas numeral_1_eq_1 = numeral_One subsection {* Setting up simplification procedures *} lemmas int_arith_rules = - neg_le_iff_le numeral_0_eq_0 numeral_1_eq_1 + neg_le_iff_le numeral_One minus_zero diff_minus left_minus right_minus mult_zero_left mult_zero_right mult_1_left mult_1_right mult_minus_left mult_minus_right @@ -1513,56 +831,39 @@ of_nat_0 of_nat_1 of_nat_Suc of_nat_add of_nat_mult of_int_0 of_int_1 of_int_add of_int_mult +use "Tools/numeral.ML" use "Tools/int_arith.ML" declaration {* K Int_Arith.setup *} -simproc_setup fast_arith ("(m::'a::{linordered_idom,number_ring}) < n" | - "(m::'a::{linordered_idom,number_ring}) <= n" | - "(m::'a::{linordered_idom,number_ring}) = n") = +simproc_setup fast_arith ("(m::'a::linordered_idom) < n" | + "(m::'a::linordered_idom) <= n" | + "(m::'a::linordered_idom) = n") = {* fn _ => fn ss => fn ct => Lin_Arith.simproc ss (term_of ct) *} setup {* Reorient_Proc.add - (fn Const (@{const_name number_of}, _) $ _ => true | _ => false) + (fn Const (@{const_name numeral}, _) $ _ => true + | Const (@{const_name neg_numeral}, _) $ _ => true + | _ => false) *} -simproc_setup reorient_numeral ("number_of w = x") = Reorient_Proc.proc +simproc_setup reorient_numeral + ("numeral w = x" | "neg_numeral w = y") = Reorient_Proc.proc subsection{*Lemmas About Small Numerals*} -lemma of_int_m1 [simp]: "of_int -1 = (-1 :: 'a :: number_ring)" -proof - - have "(of_int -1 :: 'a) = of_int (- 1)" by simp - also have "... = - of_int 1" by (simp only: of_int_minus) - also have "... = -1" by simp - finally show ?thesis . -qed - -lemma abs_minus_one [simp]: "abs (-1) = (1::'a::{linordered_idom,number_ring})" -by (simp add: abs_if) - lemma abs_power_minus_one [simp]: - "abs(-1 ^ n) = (1::'a::{linordered_idom,number_ring})" + "abs(-1 ^ n) = (1::'a::linordered_idom)" by (simp add: power_abs) -lemma of_int_number_of_eq [simp]: - "of_int (number_of v) = (number_of v :: 'a :: number_ring)" -by (simp add: number_of_eq) - text{*Lemmas for specialist use, NOT as default simprules*} (* TODO: see if semiring duplication can be removed without breaking proofs *) -lemma semiring_mult_2: "2 * z = (z+z::'a::number_semiring)" -unfolding semiring_one_add_one_is_two [symmetric] left_distrib by simp - -lemma semiring_mult_2_right: "z * 2 = (z+z::'a::number_semiring)" -by (subst mult_commute, rule semiring_mult_2) +lemma mult_2: "2 * z = (z+z::'a::semiring_1)" +unfolding one_add_one [symmetric] left_distrib by simp -lemma mult_2: "2 * z = (z+z::'a::number_ring)" -by (rule semiring_mult_2) - -lemma mult_2_right: "z * 2 = (z+z::'a::number_ring)" -by (rule semiring_mult_2_right) +lemma mult_2_right: "z * 2 = (z+z::'a::semiring_1)" +unfolding one_add_one [symmetric] right_distrib by simp subsection{*More Inequality Reasoning*} @@ -1608,7 +909,7 @@ text{*This simplifies expressions of the form @{term "int n = z"} where z is an integer literal.*} -lemmas int_eq_iff_number_of [simp] = int_eq_iff [of _ "number_of v"] for v +lemmas int_eq_iff_numeral [simp] = int_eq_iff [of _ "numeral v"] for v lemma split_nat [arith_split]: "P(nat(i::int)) = ((\n. i = int n \ P n) & (i < 0 \ P 0))" @@ -1853,12 +1154,14 @@ by (simp add: mn) finally have "2*\n\ \ 1" . thus "False" using 0 - by auto + by arith qed thus ?thesis using 0 by auto qed +ML_val {* @{const_name neg_numeral} *} + lemma pos_zmult_eq_1_iff_lemma: "(m * n = 1) ==> m = (1::int) | m = -1" by (insert abs_zmult_eq_1 [of m n], arith) @@ -1894,125 +1197,170 @@ text{*These distributive laws move literals inside sums and differences.*} -lemmas left_distrib_number_of [simp] = left_distrib [of _ _ "number_of v"] for v -lemmas right_distrib_number_of [simp] = right_distrib [of "number_of v"] for v -lemmas left_diff_distrib_number_of [simp] = left_diff_distrib [of _ _ "number_of v"] for v -lemmas right_diff_distrib_number_of [simp] = right_diff_distrib [of "number_of v"] for v +lemmas left_distrib_numeral [simp] = left_distrib [of _ _ "numeral v"] for v +lemmas right_distrib_numeral [simp] = right_distrib [of "numeral v"] for v +lemmas left_diff_distrib_numeral [simp] = left_diff_distrib [of _ _ "numeral v"] for v +lemmas right_diff_distrib_numeral [simp] = right_diff_distrib [of "numeral v"] for v text{*These are actually for fields, like real: but where else to put them?*} -lemmas zero_less_divide_iff_number_of [simp, no_atp] = zero_less_divide_iff [of "number_of w"] for w -lemmas divide_less_0_iff_number_of [simp, no_atp] = divide_less_0_iff [of "number_of w"] for w -lemmas zero_le_divide_iff_number_of [simp, no_atp] = zero_le_divide_iff [of "number_of w"] for w -lemmas divide_le_0_iff_number_of [simp, no_atp] = divide_le_0_iff [of "number_of w"] for w +lemmas zero_less_divide_iff_numeral [simp, no_atp] = zero_less_divide_iff [of "numeral w"] for w +lemmas divide_less_0_iff_numeral [simp, no_atp] = divide_less_0_iff [of "numeral w"] for w +lemmas zero_le_divide_iff_numeral [simp, no_atp] = zero_le_divide_iff [of "numeral w"] for w +lemmas divide_le_0_iff_numeral [simp, no_atp] = divide_le_0_iff [of "numeral w"] for w text {*Replaces @{text "inverse #nn"} by @{text "1/#nn"}. It looks strange, but then other simprocs simplify the quotient.*} -lemmas inverse_eq_divide_number_of [simp] = inverse_eq_divide [of "number_of w"] for w +lemmas inverse_eq_divide_numeral [simp] = + inverse_eq_divide [of "numeral w"] for w + +lemmas inverse_eq_divide_neg_numeral [simp] = + inverse_eq_divide [of "neg_numeral w"] for w text {*These laws simplify inequalities, moving unary minus from a term into the literal.*} -lemmas less_minus_iff_number_of [simp, no_atp] = less_minus_iff [of "number_of v"] for v -lemmas le_minus_iff_number_of [simp, no_atp] = le_minus_iff [of "number_of v"] for v -lemmas equation_minus_iff_number_of [simp, no_atp] = equation_minus_iff [of "number_of v"] for v -lemmas minus_less_iff_number_of [simp, no_atp] = minus_less_iff [of _ "number_of v"] for v -lemmas minus_le_iff_number_of [simp, no_atp] = minus_le_iff [of _ "number_of v"] for v -lemmas minus_equation_iff_number_of [simp, no_atp] = minus_equation_iff [of _ "number_of v"] for v +lemmas le_minus_iff_numeral [simp, no_atp] = + le_minus_iff [of "numeral v"] + le_minus_iff [of "neg_numeral v"] for v + +lemmas equation_minus_iff_numeral [simp, no_atp] = + equation_minus_iff [of "numeral v"] + equation_minus_iff [of "neg_numeral v"] for v + +lemmas minus_less_iff_numeral [simp, no_atp] = + minus_less_iff [of _ "numeral v"] + minus_less_iff [of _ "neg_numeral v"] for v + +lemmas minus_le_iff_numeral [simp, no_atp] = + minus_le_iff [of _ "numeral v"] + minus_le_iff [of _ "neg_numeral v"] for v + +lemmas minus_equation_iff_numeral [simp, no_atp] = + minus_equation_iff [of _ "numeral v"] + minus_equation_iff [of _ "neg_numeral v"] for v text{*To Simplify Inequalities Where One Side is the Constant 1*} lemma less_minus_iff_1 [simp,no_atp]: - fixes b::"'b::{linordered_idom,number_ring}" + fixes b::"'b::linordered_idom" shows "(1 < - b) = (b < -1)" by auto lemma le_minus_iff_1 [simp,no_atp]: - fixes b::"'b::{linordered_idom,number_ring}" + fixes b::"'b::linordered_idom" shows "(1 \ - b) = (b \ -1)" by auto lemma equation_minus_iff_1 [simp,no_atp]: - fixes b::"'b::number_ring" + fixes b::"'b::ring_1" shows "(1 = - b) = (b = -1)" by (subst equation_minus_iff, auto) lemma minus_less_iff_1 [simp,no_atp]: - fixes a::"'b::{linordered_idom,number_ring}" + fixes a::"'b::linordered_idom" shows "(- a < 1) = (-1 < a)" by auto lemma minus_le_iff_1 [simp,no_atp]: - fixes a::"'b::{linordered_idom,number_ring}" + fixes a::"'b::linordered_idom" shows "(- a \ 1) = (-1 \ a)" by auto lemma minus_equation_iff_1 [simp,no_atp]: - fixes a::"'b::number_ring" + fixes a::"'b::ring_1" shows "(- a = 1) = (a = -1)" by (subst minus_equation_iff, auto) text {*Cancellation of constant factors in comparisons (@{text "<"} and @{text "\"}) *} -lemmas mult_less_cancel_left_number_of [simp, no_atp] = mult_less_cancel_left [of "number_of v"] for v -lemmas mult_less_cancel_right_number_of [simp, no_atp] = mult_less_cancel_right [of _ "number_of v"] for v -lemmas mult_le_cancel_left_number_of [simp, no_atp] = mult_le_cancel_left [of "number_of v"] for v -lemmas mult_le_cancel_right_number_of [simp, no_atp] = mult_le_cancel_right [of _ "number_of v"] for v +lemmas mult_less_cancel_left_numeral [simp, no_atp] = mult_less_cancel_left [of "numeral v"] for v +lemmas mult_less_cancel_right_numeral [simp, no_atp] = mult_less_cancel_right [of _ "numeral v"] for v +lemmas mult_le_cancel_left_numeral [simp, no_atp] = mult_le_cancel_left [of "numeral v"] for v +lemmas mult_le_cancel_right_numeral [simp, no_atp] = mult_le_cancel_right [of _ "numeral v"] for v text {*Multiplying out constant divisors in comparisons (@{text "<"}, @{text "\"} and @{text "="}) *} -lemmas le_divide_eq_number_of1 [simp] = le_divide_eq [of _ _ "number_of w"] for w -lemmas divide_le_eq_number_of1 [simp] = divide_le_eq [of _ "number_of w"] for w -lemmas less_divide_eq_number_of1 [simp] = less_divide_eq [of _ _ "number_of w"] for w -lemmas divide_less_eq_number_of1 [simp] = divide_less_eq [of _ "number_of w"] for w -lemmas eq_divide_eq_number_of1 [simp] = eq_divide_eq [of _ _ "number_of w"] for w -lemmas divide_eq_eq_number_of1 [simp] = divide_eq_eq [of _ "number_of w"] for w +lemmas le_divide_eq_numeral1 [simp] = + pos_le_divide_eq [of "numeral w", OF zero_less_numeral] + neg_le_divide_eq [of "neg_numeral w", OF neg_numeral_less_zero] for w + +lemmas divide_le_eq_numeral1 [simp] = + pos_divide_le_eq [of "numeral w", OF zero_less_numeral] + neg_divide_le_eq [of "neg_numeral w", OF neg_numeral_less_zero] for w + +lemmas less_divide_eq_numeral1 [simp] = + pos_less_divide_eq [of "numeral w", OF zero_less_numeral] + neg_less_divide_eq [of "neg_numeral w", OF neg_numeral_less_zero] for w +lemmas divide_less_eq_numeral1 [simp] = + pos_divide_less_eq [of "numeral w", OF zero_less_numeral] + neg_divide_less_eq [of "neg_numeral w", OF neg_numeral_less_zero] for w + +lemmas eq_divide_eq_numeral1 [simp] = + eq_divide_eq [of _ _ "numeral w"] + eq_divide_eq [of _ _ "neg_numeral w"] for w + +lemmas divide_eq_eq_numeral1 [simp] = + divide_eq_eq [of _ "numeral w"] + divide_eq_eq [of _ "neg_numeral w"] for w subsubsection{*Optional Simplification Rules Involving Constants*} text{*Simplify quotients that are compared with a literal constant.*} -lemmas le_divide_eq_number_of = le_divide_eq [of "number_of w"] for w -lemmas divide_le_eq_number_of = divide_le_eq [of _ _ "number_of w"] for w -lemmas less_divide_eq_number_of = less_divide_eq [of "number_of w"] for w -lemmas divide_less_eq_number_of = divide_less_eq [of _ _ "number_of w"] for w -lemmas eq_divide_eq_number_of = eq_divide_eq [of "number_of w"] for w -lemmas divide_eq_eq_number_of = divide_eq_eq [of _ _ "number_of w"] for w +lemmas le_divide_eq_numeral = + le_divide_eq [of "numeral w"] + le_divide_eq [of "neg_numeral w"] for w + +lemmas divide_le_eq_numeral = + divide_le_eq [of _ _ "numeral w"] + divide_le_eq [of _ _ "neg_numeral w"] for w + +lemmas less_divide_eq_numeral = + less_divide_eq [of "numeral w"] + less_divide_eq [of "neg_numeral w"] for w + +lemmas divide_less_eq_numeral = + divide_less_eq [of _ _ "numeral w"] + divide_less_eq [of _ _ "neg_numeral w"] for w + +lemmas eq_divide_eq_numeral = + eq_divide_eq [of "numeral w"] + eq_divide_eq [of "neg_numeral w"] for w + +lemmas divide_eq_eq_numeral = + divide_eq_eq [of _ _ "numeral w"] + divide_eq_eq [of _ _ "neg_numeral w"] for w text{*Not good as automatic simprules because they cause case splits.*} lemmas divide_const_simps = - le_divide_eq_number_of divide_le_eq_number_of less_divide_eq_number_of - divide_less_eq_number_of eq_divide_eq_number_of divide_eq_eq_number_of + le_divide_eq_numeral divide_le_eq_numeral less_divide_eq_numeral + divide_less_eq_numeral eq_divide_eq_numeral divide_eq_eq_numeral le_divide_eq_1 divide_le_eq_1 less_divide_eq_1 divide_less_eq_1 text{*Division By @{text "-1"}*} -lemma divide_minus1 [simp]: - "x/-1 = -(x::'a::{field_inverse_zero, number_ring})" -by simp +lemma divide_minus1 [simp]: "(x::'a::field) / -1 = - x" + unfolding minus_one [symmetric] + unfolding nonzero_minus_divide_right [OF one_neq_zero, symmetric] + by simp -lemma minus1_divide [simp]: - "-1 / (x::'a::{field_inverse_zero, number_ring}) = - (1/x)" -by (simp add: divide_inverse) +lemma minus1_divide [simp]: "-1 / (x::'a::field) = - (1 / x)" + unfolding minus_one [symmetric] by (rule divide_minus_left) lemma half_gt_zero_iff: - "(0 < r/2) = (0 < (r::'a::{linordered_field_inverse_zero,number_ring}))" + "(0 < r/2) = (0 < (r::'a::linordered_field_inverse_zero))" by auto lemmas half_gt_zero [simp] = half_gt_zero_iff [THEN iffD2] -lemma divide_Numeral1: - "(x::'a::{field, number_ring}) / Numeral1 = x" - by simp - -lemma divide_Numeral0: - "(x::'a::{field_inverse_zero, number_ring}) / Numeral0 = 0" +lemma divide_Numeral1: "(x::'a::field) / Numeral1 = x" by simp @@ -2211,128 +1559,154 @@ subsection {* Configuration of the code generator *} -code_datatype Pls Min Bit0 Bit1 "number_of \ int \ int" +text {* Constructors *} + +definition Pos :: "num \ int" where + [simp, code_abbrev]: "Pos = numeral" + +definition Neg :: "num \ int" where + [simp, code_abbrev]: "Neg = neg_numeral" + +code_datatype "0::int" Pos Neg + + +text {* Auxiliary operations *} + +definition dup :: "int \ int" where + [simp]: "dup k = k + k" -lemmas pred_succ_numeral_code [code] = - pred_bin_simps succ_bin_simps +lemma dup_code [code]: + "dup 0 = 0" + "dup (Pos n) = Pos (Num.Bit0 n)" + "dup (Neg n) = Neg (Num.Bit0 n)" + unfolding Pos_def Neg_def neg_numeral_def + by (simp_all add: numeral_Bit0) + +definition sub :: "num \ num \ int" where + [simp]: "sub m n = numeral m - numeral n" -lemmas plus_numeral_code [code] = - add_bin_simps - arith_extra_simps(1) [where 'a = int] +lemma sub_code [code]: + "sub Num.One Num.One = 0" + "sub (Num.Bit0 m) Num.One = Pos (Num.BitM m)" + "sub (Num.Bit1 m) Num.One = Pos (Num.Bit0 m)" + "sub Num.One (Num.Bit0 n) = Neg (Num.BitM n)" + "sub Num.One (Num.Bit1 n) = Neg (Num.Bit0 n)" + "sub (Num.Bit0 m) (Num.Bit0 n) = dup (sub m n)" + "sub (Num.Bit1 m) (Num.Bit1 n) = dup (sub m n)" + "sub (Num.Bit1 m) (Num.Bit0 n) = dup (sub m n) + 1" + "sub (Num.Bit0 m) (Num.Bit1 n) = dup (sub m n) - 1" + unfolding sub_def dup_def numeral.simps Pos_def Neg_def + neg_numeral_def numeral_BitM + by (simp_all only: algebra_simps) -lemmas minus_numeral_code [code] = - minus_bin_simps - arith_extra_simps(2) [where 'a = int] - arith_extra_simps(5) [where 'a = int] + +text {* Implementations *} + +lemma one_int_code [code, code_unfold]: + "1 = Pos Num.One" + by simp + +lemma plus_int_code [code]: + "k + 0 = (k::int)" + "0 + l = (l::int)" + "Pos m + Pos n = Pos (m + n)" + "Pos m + Neg n = sub m n" + "Neg m + Pos n = sub n m" + "Neg m + Neg n = Neg (m + n)" + by simp_all -lemmas times_numeral_code [code] = - mult_bin_simps - arith_extra_simps(4) [where 'a = int] +lemma uminus_int_code [code]: + "uminus 0 = (0::int)" + "uminus (Pos m) = Neg m" + "uminus (Neg m) = Pos m" + by simp_all + +lemma minus_int_code [code]: + "k - 0 = (k::int)" + "0 - l = uminus (l::int)" + "Pos m - Pos n = sub m n" + "Pos m - Neg n = Pos (m + n)" + "Neg m - Pos n = Neg (m + n)" + "Neg m - Neg n = sub n m" + by simp_all + +lemma times_int_code [code]: + "k * 0 = (0::int)" + "0 * l = (0::int)" + "Pos m * Pos n = Pos (m * n)" + "Pos m * Neg n = Neg (m * n)" + "Neg m * Pos n = Neg (m * n)" + "Neg m * Neg n = Pos (m * n)" + by simp_all instantiation int :: equal begin definition - "HOL.equal k l \ k - l = (0\int)" + "HOL.equal k l \ k = (l::int)" -instance by default (simp add: equal_int_def) +instance by default (rule equal_int_def) end -lemma eq_number_of_int_code [code]: - "HOL.equal (number_of k \ int) (number_of l) \ HOL.equal k l" - unfolding equal_int_def number_of_is_id .. +lemma equal_int_code [code]: + "HOL.equal 0 (0::int) \ True" + "HOL.equal 0 (Pos l) \ False" + "HOL.equal 0 (Neg l) \ False" + "HOL.equal (Pos k) 0 \ False" + "HOL.equal (Pos k) (Pos l) \ HOL.equal k l" + "HOL.equal (Pos k) (Neg l) \ False" + "HOL.equal (Neg k) 0 \ False" + "HOL.equal (Neg k) (Pos l) \ False" + "HOL.equal (Neg k) (Neg l) \ HOL.equal k l" + by (auto simp add: equal) -lemma eq_int_code [code]: - "HOL.equal Int.Pls Int.Pls \ True" - "HOL.equal Int.Pls Int.Min \ False" - "HOL.equal Int.Pls (Int.Bit0 k2) \ HOL.equal Int.Pls k2" - "HOL.equal Int.Pls (Int.Bit1 k2) \ False" - "HOL.equal Int.Min Int.Pls \ False" - "HOL.equal Int.Min Int.Min \ True" - "HOL.equal Int.Min (Int.Bit0 k2) \ False" - "HOL.equal Int.Min (Int.Bit1 k2) \ HOL.equal Int.Min k2" - "HOL.equal (Int.Bit0 k1) Int.Pls \ HOL.equal k1 Int.Pls" - "HOL.equal (Int.Bit1 k1) Int.Pls \ False" - "HOL.equal (Int.Bit0 k1) Int.Min \ False" - "HOL.equal (Int.Bit1 k1) Int.Min \ HOL.equal k1 Int.Min" - "HOL.equal (Int.Bit0 k1) (Int.Bit0 k2) \ HOL.equal k1 k2" - "HOL.equal (Int.Bit0 k1) (Int.Bit1 k2) \ False" - "HOL.equal (Int.Bit1 k1) (Int.Bit0 k2) \ False" - "HOL.equal (Int.Bit1 k1) (Int.Bit1 k2) \ HOL.equal k1 k2" - unfolding equal_eq by simp_all - -lemma eq_int_refl [code nbe]: +lemma equal_int_refl [code nbe]: "HOL.equal (k::int) k \ True" - by (rule equal_refl) - -lemma less_eq_number_of_int_code [code]: - "(number_of k \ int) \ number_of l \ k \ l" - unfolding number_of_is_id .. + by (fact equal_refl) lemma less_eq_int_code [code]: - "Int.Pls \ Int.Pls \ True" - "Int.Pls \ Int.Min \ False" - "Int.Pls \ Int.Bit0 k \ Int.Pls \ k" - "Int.Pls \ Int.Bit1 k \ Int.Pls \ k" - "Int.Min \ Int.Pls \ True" - "Int.Min \ Int.Min \ True" - "Int.Min \ Int.Bit0 k \ Int.Min < k" - "Int.Min \ Int.Bit1 k \ Int.Min \ k" - "Int.Bit0 k \ Int.Pls \ k \ Int.Pls" - "Int.Bit1 k \ Int.Pls \ k < Int.Pls" - "Int.Bit0 k \ Int.Min \ k \ Int.Min" - "Int.Bit1 k \ Int.Min \ k \ Int.Min" - "Int.Bit0 k1 \ Int.Bit0 k2 \ k1 \ k2" - "Int.Bit0 k1 \ Int.Bit1 k2 \ k1 \ k2" - "Int.Bit1 k1 \ Int.Bit0 k2 \ k1 < k2" - "Int.Bit1 k1 \ Int.Bit1 k2 \ k1 \ k2" + "0 \ (0::int) \ True" + "0 \ Pos l \ True" + "0 \ Neg l \ False" + "Pos k \ 0 \ False" + "Pos k \ Pos l \ k \ l" + "Pos k \ Neg l \ False" + "Neg k \ 0 \ True" + "Neg k \ Pos l \ True" + "Neg k \ Neg l \ l \ k" by simp_all -lemma less_number_of_int_code [code]: - "(number_of k \ int) < number_of l \ k < l" - unfolding number_of_is_id .. - lemma less_int_code [code]: - "Int.Pls < Int.Pls \ False" - "Int.Pls < Int.Min \ False" - "Int.Pls < Int.Bit0 k \ Int.Pls < k" - "Int.Pls < Int.Bit1 k \ Int.Pls \ k" - "Int.Min < Int.Pls \ True" - "Int.Min < Int.Min \ False" - "Int.Min < Int.Bit0 k \ Int.Min < k" - "Int.Min < Int.Bit1 k \ Int.Min < k" - "Int.Bit0 k < Int.Pls \ k < Int.Pls" - "Int.Bit1 k < Int.Pls \ k < Int.Pls" - "Int.Bit0 k < Int.Min \ k \ Int.Min" - "Int.Bit1 k < Int.Min \ k < Int.Min" - "Int.Bit0 k1 < Int.Bit0 k2 \ k1 < k2" - "Int.Bit0 k1 < Int.Bit1 k2 \ k1 \ k2" - "Int.Bit1 k1 < Int.Bit0 k2 \ k1 < k2" - "Int.Bit1 k1 < Int.Bit1 k2 \ k1 < k2" + "0 < (0::int) \ False" + "0 < Pos l \ True" + "0 < Neg l \ False" + "Pos k < 0 \ False" + "Pos k < Pos l \ k < l" + "Pos k < Neg l \ False" + "Neg k < 0 \ True" + "Neg k < Pos l \ True" + "Neg k < Neg l \ l < k" by simp_all -definition - nat_aux :: "int \ nat \ nat" where - "nat_aux i n = nat i + n" - -lemma [code]: - "nat_aux i n = (if i \ 0 then n else nat_aux (i - 1) (Suc n))" -- {* tail recursive *} - by (auto simp add: nat_aux_def nat_eq_iff linorder_not_le order_less_imp_le - dest: zless_imp_add1_zle) +lemma nat_numeral [simp, code_abbrev]: + "nat (numeral k) = numeral k" + by (simp add: nat_eq_iff) -lemma [code]: "nat i = nat_aux i 0" - by (simp add: nat_aux_def) - -hide_const (open) nat_aux +lemma nat_code [code]: + "nat (Int.Neg k) = 0" + "nat 0 = 0" + "nat (Int.Pos k) = nat_of_num k" + by (simp_all add: nat_of_num_numeral nat_numeral) -lemma zero_is_num_zero [code, code_unfold]: - "(0\int) = Numeral0" - by simp +lemma (in ring_1) of_int_code [code]: + "of_int (Int.Neg k) = neg_numeral k" + "of_int 0 = 0" + "of_int (Int.Pos k) = numeral k" + by simp_all -lemma one_is_num_one [code, code_unfold]: - "(1\int) = Numeral1" - by simp + +text {* Serializer setup *} code_modulename SML Int Arith @@ -2345,7 +1719,7 @@ quickcheck_params [default_type = int] -hide_const (open) Pls Min Bit0 Bit1 succ pred +hide_const (open) Pos Neg sub dup subsection {* Legacy theorems *} @@ -2378,3 +1752,4 @@ lemmas zpower_int = int_power [symmetric] end + diff -r 29e92b644d6c -r f067afe98049 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/IsaMakefile Mon Mar 26 10:56:56 2012 +0200 @@ -195,6 +195,7 @@ Meson.thy \ Metis.thy \ Nat.thy \ + Num.thy \ Option.thy \ Orderings.thy \ Partial_Function.thy \ @@ -341,7 +342,6 @@ Tools/Nitpick/nitpick_util.ML \ Tools/numeral.ML \ Tools/numeral_simprocs.ML \ - Tools/numeral_syntax.ML \ Tools/Predicate_Compile/core_data.ML \ Tools/Predicate_Compile/mode_inference.ML \ Tools/Predicate_Compile/predicate_compile_aux.ML \ @@ -444,24 +444,25 @@ Library/Bit.thy Library/Boolean_Algebra.thy Library/Cardinality.thy \ Library/Char_nat.thy Library/Code_Char.thy Library/Code_Char_chr.thy \ Library/Code_Char_ord.thy Library/Code_Integer.thy \ - Library/Code_Natural.thy Library/Code_Prolog.thy \ + Library/Code_Nat.thy Library/Code_Natural.thy \ + Library/Efficient_Nat.thy Library/Code_Prolog.thy \ Library/Code_Real_Approx_By_Float.thy \ Tools/Predicate_Compile/code_prolog.ML Library/ContNotDenum.thy \ Library/Cset.thy Library/Cset_Monad.thy Library/Continuity.thy \ Library/Convex.thy Library/Countable.thy \ + Library/Dlist.thy Library/Dlist_Cset.thy Library/Eval_Witness.thy \ Library/DAList.thy Library/Dlist.thy Library/Dlist_Cset.thy \ - Library/Efficient_Nat.thy Library/Eval_Witness.thy \ + Library/Eval_Witness.thy \ Library/Extended_Real.thy Library/Extended_Nat.thy Library/Float.thy \ Library/Formal_Power_Series.thy Library/Fraction_Field.thy \ Library/FrechetDeriv.thy Library/Cset.thy Library/FuncSet.thy \ - Library/Function_Algebras.thy \ - Library/Fundamental_Theorem_Algebra.thy Library/Glbs.thy \ - Library/Indicator_Function.thy Library/Infinite_Set.thy \ - Library/Inner_Product.thy Library/Kleene_Algebra.thy \ - Library/LaTeXsugar.thy Library/Lattice_Algebras.thy \ - Library/Lattice_Syntax.thy Library/Library.thy Library/List_Cset.thy \ - Library/List_Prefix.thy Library/List_lexord.thy Library/Mapping.thy \ - Library/Monad_Syntax.thy \ + Library/Function_Algebras.thy Library/Fundamental_Theorem_Algebra.thy \ + Library/Glbs.thy Library/Indicator_Function.thy \ + Library/Infinite_Set.thy Library/Inner_Product.thy \ + Library/Kleene_Algebra.thy Library/LaTeXsugar.thy \ + Library/Lattice_Algebras.thy Library/Lattice_Syntax.thy \ + Library/Library.thy Library/List_Cset.thy Library/List_Prefix.thy \ + Library/List_lexord.thy Library/Mapping.thy Library/Monad_Syntax.thy \ Library/Multiset.thy Library/Nat_Bijection.thy \ Library/Numeral_Type.thy Library/Old_Recdef.thy \ Library/OptionalSugar.thy Library/Order_Relation.thy \ @@ -479,7 +480,7 @@ Library/State_Monad.thy Library/Ramsey.thy \ Library/Reflection.thy Library/Sublist_Order.thy \ Library/Sum_of_Squares.thy Library/Sum_of_Squares/sos_wrapper.ML \ - Library/Sum_of_Squares/sum_of_squares.ML \ + Library/Sum_of_Squares/sum_of_squares.ML Library/Target_Numeral.thy \ Library/Transitive_Closure_Table.thy Library/Univ_Poly.thy \ Library/Wfrec.thy Library/While_Combinator.thy Library/Zorn.thy \ $(SRC)/Tools/adhoc_overloading.ML Library/positivstellensatz.ML \ @@ -758,11 +759,11 @@ HOL-Library-Codegenerator_Test: HOL-Library $(LOG)/HOL-Library-Codegenerator_Test.gz -$(LOG)/HOL-Library-Codegenerator_Test.gz: $(OUT)/HOL-Library \ - Codegenerator_Test/ROOT.ML \ - Codegenerator_Test/Candidates.thy \ - Codegenerator_Test/Candidates_Pretty.thy \ - Codegenerator_Test/Generate.thy \ +$(LOG)/HOL-Library-Codegenerator_Test.gz: $(OUT)/HOL-Library \ + Codegenerator_Test/ROOT.ML \ + Codegenerator_Test/Candidates.thy \ + Codegenerator_Test/Candidates_Pretty.thy \ + Codegenerator_Test/Generate.thy \ Codegenerator_Test/Generate_Pretty.thy @$(ISABELLE_TOOL) usedir -d false -g false -i false $(OUT)/HOL-Library Codegenerator_Test @@ -920,6 +921,10 @@ HOL-Imperative_HOL: HOL $(LOG)/HOL-Imperative_HOL.gz $(LOG)/HOL-Imperative_HOL.gz: $(OUT)/HOL \ + Library/Code_Integer.thy \ + Library/Code_Nat.thy \ + Library/Code_Natural.thy \ + Library/Efficient_Nat.thy \ Imperative_HOL/Array.thy \ Imperative_HOL/Heap.thy \ Imperative_HOL/Heap_Monad.thy \ @@ -943,6 +948,10 @@ HOL-Decision_Procs: HOL $(LOG)/HOL-Decision_Procs.gz $(LOG)/HOL-Decision_Procs.gz: $(OUT)/HOL \ + Library/Code_Integer.thy \ + Library/Code_Nat.thy \ + Library/Code_Natural.thy \ + Library/Efficient_Nat.thy \ Decision_Procs/Approximation.thy \ Decision_Procs/Commutative_Ring.thy \ Decision_Procs/Commutative_Ring_Complete.thy \ @@ -991,9 +1000,12 @@ HOL-Proofs-Extraction: HOL-Proofs $(LOG)/HOL-Proofs-Extraction.gz $(LOG)/HOL-Proofs-Extraction.gz: $(OUT)/HOL-Proofs \ - Library/Efficient_Nat.thy Proofs/Extraction/Euclid.thy \ + Library/Code_Integer.thy Library/Code_Nat.thy \ + Library/Code_Natural.thy Library/Efficient_Nat.thy \ + Proofs/Extraction/Euclid.thy \ Proofs/Extraction/Greatest_Common_Divisor.thy \ - Proofs/Extraction/Higman.thy Proofs/Extraction/Higman_Extraction.thy \ + Proofs/Extraction/Higman.thy \ + Proofs/Extraction/Higman_Extraction.thy \ Proofs/Extraction/Pigeonhole.thy \ Proofs/Extraction/QuotRem.thy Proofs/Extraction/ROOT.ML \ Proofs/Extraction/Util.thy Proofs/Extraction/Warshall.thy \ @@ -1113,15 +1125,17 @@ HOL-ex: HOL $(LOG)/HOL-ex.gz $(LOG)/HOL-ex.gz: $(OUT)/HOL Decision_Procs/Commutative_Ring.thy \ + Library/Code_Integer.thy Library/Code_Nat.thy \ + Library/Code_Natural.thy Library/Efficient_Nat.thy \ Number_Theory/Primes.thy ex/Abstract_NAT.thy ex/Antiquote.thy \ ex/Arith_Examples.thy ex/Arithmetic_Series_Complex.thy ex/BT.thy \ ex/BinEx.thy ex/Binary.thy ex/Birthday_Paradox.thy ex/CTL.thy \ ex/Case_Product.thy ex/Chinese.thy ex/Classical.thy \ - ex/Coercion_Examples.thy ex/Coherent.thy \ - ex/Dedekind_Real.thy ex/Efficient_Nat_examples.thy \ + ex/Code_Nat_examples.thy \ + ex/Coercion_Examples.thy ex/Coherent.thy ex/Dedekind_Real.thy \ ex/Eval_Examples.thy ex/Executable_Relation.thy ex/Fundefs.thy \ ex/Gauge_Integration.thy ex/Groebner_Examples.thy ex/Guess.thy \ - ex/HarmonicSeries.thy ex/Hebrew.thy ex/Hex_Bin_Examples.thy \ + ex/HarmonicSeries.thy ex/Hebrew.thy ex/Hex_Bin_Examples.thy \ ex/Higher_Order_Logic.thy ex/Iff_Oracle.thy ex/Induction_Schema.thy \ ex/Interpretation_with_Defs.thy ex/Intuitionistic.thy \ ex/Lagrange.thy ex/List_to_Set_Comprehension_Examples.thy \ diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/BigO.thy --- a/src/HOL/Library/BigO.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/BigO.thy Mon Mar 26 10:56:56 2012 +0200 @@ -132,7 +132,6 @@ apply (simp add: abs_triangle_ineq) apply (simp add: order_less_le) apply (rule mult_nonneg_nonneg) - apply (rule add_nonneg_nonneg) apply auto apply (rule_tac x = "%n. if (abs (f n)) < abs (g n) then x n else 0" in exI) @@ -150,11 +149,8 @@ apply (rule abs_triangle_ineq) apply (simp add: order_less_le) apply (rule mult_nonneg_nonneg) - apply (rule add_nonneg_nonneg) - apply (erule order_less_imp_le)+ + apply (erule order_less_imp_le) apply simp - apply (rule ext) - apply (auto simp add: if_splits linorder_not_le) done lemma bigo_plus_subset2 [intro]: "A <= O(f) ==> B <= O(f) ==> A \ B <= O(f)" diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Binomial.thy --- a/src/HOL/Library/Binomial.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Binomial.thy Mon Mar 26 10:56:56 2012 +0200 @@ -350,7 +350,7 @@ have eq: "(- (1\'a)) ^ n = setprod (\i. - 1) {0 .. n - 1}" by auto from n0 have ?thesis - by (simp add: pochhammer_def gbinomial_def field_simps eq setprod_timesf[symmetric])} + by (simp add: pochhammer_def gbinomial_def field_simps eq setprod_timesf[symmetric] del: minus_one) (* FIXME: del: minus_one *)} ultimately show ?thesis by blast qed @@ -417,8 +417,8 @@ from eq[symmetric] have ?thesis using kn apply (simp add: binomial_fact[OF kn, where ?'a = 'a] - gbinomial_pochhammer field_simps pochhammer_Suc_setprod) - apply (simp add: pochhammer_Suc_setprod fact_altdef_nat h of_nat_setprod setprod_timesf[symmetric] eq' del: One_nat_def power_Suc) + gbinomial_pochhammer field_simps pochhammer_Suc_setprod del: minus_one) + apply (simp add: pochhammer_Suc_setprod fact_altdef_nat h of_nat_setprod setprod_timesf[symmetric] eq' del: One_nat_def power_Suc del: minus_one) unfolding setprod_Un_disjoint[OF th0, unfolded eq3, of "of_nat:: nat \ 'a"] eq[unfolded h] unfolding mult_assoc[symmetric] unfolding setprod_timesf[symmetric] diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Bit.thy --- a/src/HOL/Library/Bit.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Bit.thy Mon Mar 26 10:56:56 2012 +0200 @@ -96,27 +96,18 @@ subsection {* Numerals at type @{typ bit} *} -instantiation bit :: number_ring -begin - -definition number_of_bit_def: - "(number_of w :: bit) = of_int w" - -instance proof -qed (rule number_of_bit_def) - -end - text {* All numerals reduce to either 0 or 1. *} lemma bit_minus1 [simp]: "-1 = (1 :: bit)" - by (simp only: number_of_Min uminus_bit_def) + by (simp only: minus_one [symmetric] uminus_bit_def) + +lemma bit_neg_numeral [simp]: "(neg_numeral w :: bit) = numeral w" + by (simp only: neg_numeral_def uminus_bit_def) -lemma bit_number_of_even [simp]: "number_of (Int.Bit0 w) = (0 :: bit)" - by (simp only: number_of_Bit0 add_0_left bit_add_self) +lemma bit_numeral_even [simp]: "numeral (Num.Bit0 w) = (0 :: bit)" + by (simp only: numeral_Bit0 bit_add_self) -lemma bit_number_of_odd [simp]: "number_of (Int.Bit1 w) = (1 :: bit)" - by (simp only: number_of_Bit1 add_assoc bit_add_self - monoid_add_class.add_0_right) +lemma bit_numeral_odd [simp]: "numeral (Num.Bit1 w) = (1 :: bit)" + by (simp only: numeral_Bit1 bit_add_self add_0_left) end diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Cardinality.thy --- a/src/HOL/Library/Cardinality.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Cardinality.thy Mon Mar 26 10:56:56 2012 +0200 @@ -5,7 +5,7 @@ header {* Cardinality of types *} theory Cardinality -imports Main +imports "~~/src/HOL/Main" begin subsection {* Preliminary lemmas *} diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Code_Integer.thy --- a/src/HOL/Library/Code_Integer.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Code_Integer.thy Mon Mar 26 10:56:56 2012 +0200 @@ -9,6 +9,43 @@ begin text {* + Representation-ignorant code equations for conversions. +*} + +lemma nat_code [code]: + "nat k = (if k \ 0 then 0 else + let + (l, j) = divmod_int k 2; + l' = 2 * nat l + in if j = 0 then l' else Suc l')" +proof - + have "2 = nat 2" by simp + show ?thesis + apply (auto simp add: Let_def divmod_int_mod_div not_le + nat_div_distrib nat_mult_distrib mult_div_cancel mod_2_not_eq_zero_eq_one_int) + apply (unfold `2 = nat 2`) + apply (subst nat_mod_distrib [symmetric]) + apply simp_all + done +qed + +lemma (in ring_1) of_int_code: + "of_int k = (if k = 0 then 0 + else if k < 0 then - of_int (- k) + else let + (l, j) = divmod_int k 2; + l' = 2 * of_int l + in if j = 0 then l' else l' + 1)" +proof - + from mod_div_equality have *: "of_int k = of_int (k div 2 * 2 + k mod 2)" by simp + show ?thesis + by (simp add: Let_def divmod_int_mod_div mod_2_not_eq_zero_eq_one_int + of_int_add [symmetric]) (simp add: * mult_commute) +qed + +declare of_int_code [code] + +text {* HOL numeral expressions are mapped to integer literals in target languages, using predefined target language operations for abstract integer operations. @@ -24,42 +61,21 @@ code_instance int :: equal (Haskell -) +code_const "0::int" + (SML "0") + (OCaml "Big'_int.zero'_big'_int") + (Haskell "0") + (Scala "BigInt(0)") + setup {* - fold (Numeral.add_code @{const_name number_int_inst.number_of_int} - true Code_Printer.literal_numeral) ["SML", "OCaml", "Haskell", "Scala"] + fold (Numeral.add_code @{const_name Int.Pos} + false Code_Printer.literal_numeral) ["SML", "OCaml", "Haskell", "Scala"] *} -code_const "Int.Pls" and "Int.Min" and "Int.Bit0" and "Int.Bit1" - (SML "raise/ Fail/ \"Pls\"" - and "raise/ Fail/ \"Min\"" - and "!((_);/ raise/ Fail/ \"Bit0\")" - and "!((_);/ raise/ Fail/ \"Bit1\")") - (OCaml "failwith/ \"Pls\"" - and "failwith/ \"Min\"" - and "!((_);/ failwith/ \"Bit0\")" - and "!((_);/ failwith/ \"Bit1\")") - (Haskell "error/ \"Pls\"" - and "error/ \"Min\"" - and "error/ \"Bit0\"" - and "error/ \"Bit1\"") - (Scala "!error(\"Pls\")" - and "!error(\"Min\")" - and "!error(\"Bit0\")" - and "!error(\"Bit1\")") - -code_const Int.pred - (SML "IntInf.- ((_), 1)") - (OCaml "Big'_int.pred'_big'_int") - (Haskell "!(_/ -/ 1)") - (Scala "!(_ -/ 1)") - (Eval "!(_/ -/ 1)") - -code_const Int.succ - (SML "IntInf.+ ((_), 1)") - (OCaml "Big'_int.succ'_big'_int") - (Haskell "!(_/ +/ 1)") - (Scala "!(_ +/ 1)") - (Eval "!(_/ +/ 1)") +setup {* + fold (Numeral.add_code @{const_name Int.Neg} + true Code_Printer.literal_numeral) ["SML", "OCaml", "Haskell", "Scala"] +*} code_const "op + \ int \ int \ int" (SML "IntInf.+ ((_), (_))") @@ -82,6 +98,19 @@ (Scala infixl 7 "-") (Eval infixl 8 "-") +code_const Int.dup + (SML "IntInf.*/ (2,/ (_))") + (OCaml "Big'_int.mult'_big'_int/ 2") + (Haskell "!(2 * _)") + (Scala "!(2 * _)") + (Eval "!(2 * _)") + +code_const Int.sub + (SML "!(raise/ Fail/ \"sub\")") + (OCaml "failwith/ \"sub\"") + (Haskell "error/ \"sub\"") + (Scala "!error(\"sub\")") + code_const "op * \ int \ int \ int" (SML "IntInf.* ((_), (_))") (OCaml "Big'_int.mult'_big'_int") @@ -124,9 +153,7 @@ (Scala "!_.as'_BigInt") (Eval "_") -text {* Evaluation *} - code_const "Code_Evaluation.term_of \ int \ term" (Eval "HOLogic.mk'_number/ HOLogic.intT") -end \ No newline at end of file +end diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Code_Nat.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Code_Nat.thy Mon Mar 26 10:56:56 2012 +0200 @@ -0,0 +1,258 @@ +(* Title: HOL/Library/Code_Nat.thy + Author: Stefan Berghofer, Florian Haftmann, TU Muenchen +*) + +header {* Implementation of natural numbers as binary numerals *} + +theory Code_Nat +imports Main +begin + +text {* + When generating code for functions on natural numbers, the + canonical representation using @{term "0::nat"} and + @{term Suc} is unsuitable for computations involving large + numbers. This theory refines the representation of + natural numbers for code generation to use binary + numerals, which do not grow linear in size but logarithmic. +*} + +subsection {* Representation *} + +lemma [code_abbrev]: + "nat_of_num = numeral" + by (fact nat_of_num_numeral) + +code_datatype "0::nat" nat_of_num + +lemma [code]: + "num_of_nat 0 = Num.One" + "num_of_nat (nat_of_num k) = k" + by (simp_all add: nat_of_num_inverse) + +lemma [code]: + "(1\nat) = Numeral1" + by simp + +lemma [code_abbrev]: "Numeral1 = (1\nat)" + by simp + +lemma [code]: + "Suc n = n + 1" + by simp + + +subsection {* Basic arithmetic *} + +lemma [code, code del]: + "(plus :: nat \ _) = plus" .. + +lemma plus_nat_code [code]: + "nat_of_num k + nat_of_num l = nat_of_num (k + l)" + "m + 0 = (m::nat)" + "0 + n = (n::nat)" + by (simp_all add: nat_of_num_numeral) + +text {* Bounded subtraction needs some auxiliary *} + +definition dup :: "nat \ nat" where + "dup n = n + n" + +lemma dup_code [code]: + "dup 0 = 0" + "dup (nat_of_num k) = nat_of_num (Num.Bit0 k)" + unfolding Num_def by (simp_all add: dup_def numeral_Bit0) + +definition sub :: "num \ num \ nat option" where + "sub k l = (if k \ l then Some (numeral k - numeral l) else None)" + +lemma sub_code [code]: + "sub Num.One Num.One = Some 0" + "sub (Num.Bit0 m) Num.One = Some (nat_of_num (Num.BitM m))" + "sub (Num.Bit1 m) Num.One = Some (nat_of_num (Num.Bit0 m))" + "sub Num.One (Num.Bit0 n) = None" + "sub Num.One (Num.Bit1 n) = None" + "sub (Num.Bit0 m) (Num.Bit0 n) = Option.map dup (sub m n)" + "sub (Num.Bit1 m) (Num.Bit1 n) = Option.map dup (sub m n)" + "sub (Num.Bit1 m) (Num.Bit0 n) = Option.map (\q. dup q + 1) (sub m n)" + "sub (Num.Bit0 m) (Num.Bit1 n) = (case sub m n of None \ None + | Some q \ if q = 0 then None else Some (dup q - 1))" + apply (auto simp add: nat_of_num_numeral + Num.dbl_def Num.dbl_inc_def Num.dbl_dec_def + Let_def le_imp_diff_is_add BitM_plus_one sub_def dup_def) + apply (simp_all add: sub_non_positive) + apply (simp_all add: sub_non_negative [symmetric, where ?'a = int]) + done + +lemma [code, code del]: + "(minus :: nat \ _) = minus" .. + +lemma minus_nat_code [code]: + "nat_of_num k - nat_of_num l = (case sub k l of None \ 0 | Some j \ j)" + "m - 0 = (m::nat)" + "0 - n = (0::nat)" + by (simp_all add: nat_of_num_numeral sub_non_positive sub_def) + +lemma [code, code del]: + "(times :: nat \ _) = times" .. + +lemma times_nat_code [code]: + "nat_of_num k * nat_of_num l = nat_of_num (k * l)" + "m * 0 = (0::nat)" + "0 * n = (0::nat)" + by (simp_all add: nat_of_num_numeral) + +lemma [code, code del]: + "(HOL.equal :: nat \ _) = HOL.equal" .. + +lemma equal_nat_code [code]: + "HOL.equal 0 (0::nat) \ True" + "HOL.equal 0 (nat_of_num l) \ False" + "HOL.equal (nat_of_num k) 0 \ False" + "HOL.equal (nat_of_num k) (nat_of_num l) \ HOL.equal k l" + by (simp_all add: nat_of_num_numeral equal) + +lemma equal_nat_refl [code nbe]: + "HOL.equal (n::nat) n \ True" + by (rule equal_refl) + +lemma [code, code del]: + "(less_eq :: nat \ _) = less_eq" .. + +lemma less_eq_nat_code [code]: + "0 \ (n::nat) \ True" + "nat_of_num k \ 0 \ False" + "nat_of_num k \ nat_of_num l \ k \ l" + by (simp_all add: nat_of_num_numeral) + +lemma [code, code del]: + "(less :: nat \ _) = less" .. + +lemma less_nat_code [code]: + "(m::nat) < 0 \ False" + "0 < nat_of_num l \ True" + "nat_of_num k < nat_of_num l \ k < l" + by (simp_all add: nat_of_num_numeral) + + +subsection {* Conversions *} + +lemma [code, code del]: + "of_nat = of_nat" .. + +lemma of_nat_code [code]: + "of_nat 0 = 0" + "of_nat (nat_of_num k) = numeral k" + by (simp_all add: nat_of_num_numeral) + + +subsection {* Case analysis *} + +text {* + Case analysis on natural numbers is rephrased using a conditional + expression: +*} + +lemma [code, code_unfold]: + "nat_case = (\f g n. if n = 0 then f else g (n - 1))" + by (auto simp add: fun_eq_iff dest!: gr0_implies_Suc) + + +subsection {* Preprocessors *} + +text {* + The term @{term "Suc n"} is no longer a valid pattern. + Therefore, all occurrences of this term in a position + where a pattern is expected (i.e.~on the left-hand side of a recursion + equation) must be eliminated. + This can be accomplished by applying the following transformation rules: +*} + +lemma Suc_if_eq: "(\n. f (Suc n) \ h n) \ f 0 \ g \ + f n \ if n = 0 then g else h (n - 1)" + by (rule eq_reflection) (cases n, simp_all) + +text {* + The rules above are built into a preprocessor that is plugged into + the code generator. Since the preprocessor for introduction rules + does not know anything about modes, some of the modes that worked + for the canonical representation of natural numbers may no longer work. +*} + +(*<*) +setup {* +let + +fun remove_suc thy thms = + let + val vname = singleton (Name.variant_list (map fst + (fold (Term.add_var_names o Thm.full_prop_of) thms []))) "n"; + val cv = cterm_of thy (Var ((vname, 0), HOLogic.natT)); + fun lhs_of th = snd (Thm.dest_comb + (fst (Thm.dest_comb (cprop_of th)))); + fun rhs_of th = snd (Thm.dest_comb (cprop_of th)); + fun find_vars ct = (case term_of ct of + (Const (@{const_name Suc}, _) $ Var _) => [(cv, snd (Thm.dest_comb ct))] + | _ $ _ => + let val (ct1, ct2) = Thm.dest_comb ct + in + map (apfst (fn ct => Thm.apply ct ct2)) (find_vars ct1) @ + map (apfst (Thm.apply ct1)) (find_vars ct2) + end + | _ => []); + val eqs = maps + (fn th => map (pair th) (find_vars (lhs_of th))) thms; + fun mk_thms (th, (ct, cv')) = + let + val th' = + Thm.implies_elim + (Conv.fconv_rule (Thm.beta_conversion true) + (Drule.instantiate' + [SOME (ctyp_of_term ct)] [SOME (Thm.lambda cv ct), + SOME (Thm.lambda cv' (rhs_of th)), NONE, SOME cv'] + @{thm Suc_if_eq})) (Thm.forall_intr cv' th) + in + case map_filter (fn th'' => + SOME (th'', singleton + (Variable.trade (K (fn [th'''] => [th''' RS th'])) + (Variable.global_thm_context th'')) th'') + handle THM _ => NONE) thms of + [] => NONE + | thps => + let val (ths1, ths2) = split_list thps + in SOME (subtract Thm.eq_thm (th :: ths1) thms @ ths2) end + end + in get_first mk_thms eqs end; + +fun eqn_suc_base_preproc thy thms = + let + val dest = fst o Logic.dest_equals o prop_of; + val contains_suc = exists_Const (fn (c, _) => c = @{const_name Suc}); + in + if forall (can dest) thms andalso exists (contains_suc o dest) thms + then thms |> perhaps_loop (remove_suc thy) |> (Option.map o map) Drule.zero_var_indexes + else NONE + end; + +val eqn_suc_preproc = Code_Preproc.simple_functrans eqn_suc_base_preproc; + +in + + Code_Preproc.add_functrans ("eqn_Suc", eqn_suc_preproc) + +end; +*} +(*>*) + +code_modulename SML + Code_Nat Arith + +code_modulename OCaml + Code_Nat Arith + +code_modulename Haskell + Code_Nat Arith + +hide_const (open) dup sub + +end diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Code_Natural.thy --- a/src/HOL/Library/Code_Natural.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Code_Natural.thy Mon Mar 26 10:56:56 2012 +0200 @@ -106,22 +106,26 @@ (Scala "Natural") setup {* - fold (Numeral.add_code @{const_name number_code_numeral_inst.number_of_code_numeral} + fold (Numeral.add_code @{const_name Code_Numeral.Num} false Code_Printer.literal_alternative_numeral) ["Haskell", "Scala"] *} code_instance code_numeral :: equal (Haskell -) -code_const "op + \ code_numeral \ code_numeral \ code_numeral" +code_const "0::code_numeral" + (Haskell "0") + (Scala "Natural(0)") + +code_const "plus \ code_numeral \ code_numeral \ code_numeral" (Haskell infixl 6 "+") (Scala infixl 7 "+") -code_const "op - \ code_numeral \ code_numeral \ code_numeral" +code_const "minus \ code_numeral \ code_numeral \ code_numeral" (Haskell infixl 6 "-") (Scala infixl 7 "-") -code_const "op * \ code_numeral \ code_numeral \ code_numeral" +code_const "times \ code_numeral \ code_numeral \ code_numeral" (Haskell infixl 7 "*") (Scala infixl 8 "*") @@ -133,11 +137,11 @@ (Haskell infix 4 "==") (Scala infixl 5 "==") -code_const "op \ \ code_numeral \ code_numeral \ bool" +code_const "less_eq \ code_numeral \ code_numeral \ bool" (Haskell infix 4 "<=") (Scala infixl 4 "<=") -code_const "op < \ code_numeral \ code_numeral \ bool" +code_const "less \ code_numeral \ code_numeral \ bool" (Haskell infix 4 "<") (Scala infixl 4 "<") diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Code_Prolog.thy --- a/src/HOL/Library/Code_Prolog.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Code_Prolog.thy Mon Mar 26 10:56:56 2012 +0200 @@ -11,8 +11,10 @@ section {* Setup for Numerals *} -setup {* Predicate_Compile_Data.ignore_consts [@{const_name number_of}] *} -setup {* Predicate_Compile_Data.keep_functions [@{const_name number_of}] *} +setup {* Predicate_Compile_Data.ignore_consts + [@{const_name numeral}, @{const_name neg_numeral}] *} + +setup {* Predicate_Compile_Data.keep_functions + [@{const_name numeral}, @{const_name neg_numeral}] *} end - diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Code_Real_Approx_By_Float.thy --- a/src/HOL/Library/Code_Real_Approx_By_Float.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Code_Real_Approx_By_Float.thy Mon Mar 26 10:56:56 2012 +0200 @@ -129,9 +129,23 @@ lemma of_int_eq_real_of_int[code_unfold]: "of_int = real_of_int" unfolding real_of_int_def .. -hide_const (open) real_of_int +lemma [code_unfold del]: + "0 \ (of_rat 0 :: real)" + by simp + +lemma [code_unfold del]: + "1 \ (of_rat 1 :: real)" + by simp -declare number_of_real_code [code_unfold del] +lemma [code_unfold del]: + "numeral k \ (of_rat (numeral k) :: real)" + by simp + +lemma [code_unfold del]: + "neg_numeral k \ (of_rat (neg_numeral k) :: real)" + by simp + +hide_const (open) real_of_int notepad begin diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Efficient_Nat.thy --- a/src/HOL/Library/Efficient_Nat.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Efficient_Nat.thy Mon Mar 26 10:56:56 2012 +0200 @@ -5,175 +5,16 @@ header {* Implementation of natural numbers by target-language integers *} theory Efficient_Nat -imports Code_Integer Main +imports Code_Nat Code_Integer Main begin text {* - When generating code for functions on natural numbers, the - canonical representation using @{term "0::nat"} and - @{term Suc} is unsuitable for computations involving large - numbers. The efficiency of the generated code can be improved + The efficiency of the generated code for natural numbers can be improved drastically by implementing natural numbers by target-language integers. To do this, just include this theory. *} -subsection {* Basic arithmetic *} - -text {* - Most standard arithmetic functions on natural numbers are implemented - using their counterparts on the integers: -*} - -code_datatype number_nat_inst.number_of_nat - -lemma zero_nat_code [code, code_unfold]: - "0 = (Numeral0 :: nat)" - by simp - -lemma one_nat_code [code, code_unfold]: - "1 = (Numeral1 :: nat)" - by simp - -lemma Suc_code [code]: - "Suc n = n + 1" - by simp - -lemma plus_nat_code [code]: - "n + m = nat (of_nat n + of_nat m)" - by simp - -lemma minus_nat_code [code]: - "n - m = nat (of_nat n - of_nat m)" - by simp - -lemma times_nat_code [code]: - "n * m = nat (of_nat n * of_nat m)" - unfolding of_nat_mult [symmetric] by simp - -lemma divmod_nat_code [code]: - "divmod_nat n m = map_pair nat nat (pdivmod (of_nat n) (of_nat m))" - by (simp add: map_pair_def split_def pdivmod_def nat_div_distrib nat_mod_distrib divmod_nat_div_mod) - -lemma eq_nat_code [code]: - "HOL.equal n m \ HOL.equal (of_nat n \ int) (of_nat m)" - by (simp add: equal) - -lemma eq_nat_refl [code nbe]: - "HOL.equal (n::nat) n \ True" - by (rule equal_refl) - -lemma less_eq_nat_code [code]: - "n \ m \ (of_nat n \ int) \ of_nat m" - by simp - -lemma less_nat_code [code]: - "n < m \ (of_nat n \ int) < of_nat m" - by simp - -subsection {* Case analysis *} - -text {* - Case analysis on natural numbers is rephrased using a conditional - expression: -*} - -lemma [code, code_unfold]: - "nat_case = (\f g n. if n = 0 then f else g (n - 1))" - by (auto simp add: fun_eq_iff dest!: gr0_implies_Suc) - - -subsection {* Preprocessors *} - -text {* - In contrast to @{term "Suc n"}, the term @{term "n + (1::nat)"} is no longer - a constructor term. Therefore, all occurrences of this term in a position - where a pattern is expected (i.e.\ on the left-hand side of a recursion - equation or in the arguments of an inductive relation in an introduction - rule) must be eliminated. - This can be accomplished by applying the following transformation rules: -*} - -lemma Suc_if_eq: "(\n. f (Suc n) \ h n) \ f 0 \ g \ - f n \ if n = 0 then g else h (n - 1)" - by (rule eq_reflection) (cases n, simp_all) - -lemma Suc_clause: "(\n. P n (Suc n)) \ n \ 0 \ P (n - 1) n" - by (cases n) simp_all - -text {* - The rules above are built into a preprocessor that is plugged into - the code generator. Since the preprocessor for introduction rules - does not know anything about modes, some of the modes that worked - for the canonical representation of natural numbers may no longer work. -*} - -(*<*) -setup {* -let - -fun remove_suc thy thms = - let - val vname = singleton (Name.variant_list (map fst - (fold (Term.add_var_names o Thm.full_prop_of) thms []))) "n"; - val cv = cterm_of thy (Var ((vname, 0), HOLogic.natT)); - fun lhs_of th = snd (Thm.dest_comb - (fst (Thm.dest_comb (cprop_of th)))); - fun rhs_of th = snd (Thm.dest_comb (cprop_of th)); - fun find_vars ct = (case term_of ct of - (Const (@{const_name Suc}, _) $ Var _) => [(cv, snd (Thm.dest_comb ct))] - | _ $ _ => - let val (ct1, ct2) = Thm.dest_comb ct - in - map (apfst (fn ct => Thm.apply ct ct2)) (find_vars ct1) @ - map (apfst (Thm.apply ct1)) (find_vars ct2) - end - | _ => []); - val eqs = maps - (fn th => map (pair th) (find_vars (lhs_of th))) thms; - fun mk_thms (th, (ct, cv')) = - let - val th' = - Thm.implies_elim - (Conv.fconv_rule (Thm.beta_conversion true) - (Drule.instantiate' - [SOME (ctyp_of_term ct)] [SOME (Thm.lambda cv ct), - SOME (Thm.lambda cv' (rhs_of th)), NONE, SOME cv'] - @{thm Suc_if_eq})) (Thm.forall_intr cv' th) - in - case map_filter (fn th'' => - SOME (th'', singleton - (Variable.trade (K (fn [th'''] => [th''' RS th'])) - (Variable.global_thm_context th'')) th'') - handle THM _ => NONE) thms of - [] => NONE - | thps => - let val (ths1, ths2) = split_list thps - in SOME (subtract Thm.eq_thm (th :: ths1) thms @ ths2) end - end - in get_first mk_thms eqs end; - -fun eqn_suc_base_preproc thy thms = - let - val dest = fst o Logic.dest_equals o prop_of; - val contains_suc = exists_Const (fn (c, _) => c = @{const_name Suc}); - in - if forall (can dest) thms andalso exists (contains_suc o dest) thms - then thms |> perhaps_loop (remove_suc thy) |> (Option.map o map) Drule.zero_var_indexes - else NONE - end; - -val eqn_suc_preproc = Code_Preproc.simple_functrans eqn_suc_base_preproc; - -in - - Code_Preproc.add_functrans ("eqn_Suc", eqn_suc_preproc) - -end; -*} -(*>*) - - -subsection {* Target language setup *} +subsection {* Target language fundamentals *} text {* For ML, we map @{typ nat} to target language integers, where we @@ -282,47 +123,32 @@ code_instance nat :: equal (Haskell -) -text {* - Natural numerals. -*} - -lemma [code_abbrev]: - "number_nat_inst.number_of_nat i = nat (number_of i)" - -- {* this interacts as desired with @{thm nat_number_of_def} *} - by (simp add: number_nat_inst.number_of_nat) - setup {* - fold (Numeral.add_code @{const_name number_nat_inst.number_of_nat} + fold (Numeral.add_code @{const_name nat_of_num} false Code_Printer.literal_positive_numeral) ["SML", "OCaml", "Haskell", "Scala"] *} +code_const "0::nat" + (SML "0") + (OCaml "Big'_int.zero'_big'_int") + (Haskell "0") + (Scala "Nat(0)") + + +subsection {* Conversions *} + text {* Since natural numbers are implemented - using integers in ML, the coercion function @{const "of_nat"} of type + using integers in ML, the coercion function @{term "int"} of type @{typ "nat \ int"} is simply implemented by the identity function. For the @{const nat} function for converting an integer to a natural - number, we give a specific implementation using an ML function that + number, we give a specific implementation using an ML expression that returns its input value, provided that it is non-negative, and otherwise returns @{text "0"}. *} definition int :: "nat \ int" where - [code del, code_abbrev]: "int = of_nat" - -lemma int_code' [code]: - "int (number_of l) = (if neg (number_of l \ int) then 0 else number_of l)" - unfolding int_nat_number_of [folded int_def] .. - -lemma nat_code' [code]: - "nat (number_of l) = (if neg (number_of l \ int) then 0 else number_of l)" - unfolding nat_number_of_def number_of_is_id neg_def by simp - -lemma of_nat_int: (* FIXME delete candidate *) - "of_nat = int" by (simp add: int_def) - -lemma of_nat_aux_int [code_unfold]: - "of_nat_aux (\i. i + 1) k 0 = int k" - by (simp add: int_def Nat.of_nat_code) + [code_abbrev]: "int = of_nat" code_const int (SML "_") @@ -331,7 +157,7 @@ code_const nat (SML "IntInf.max/ (0,/ _)") (OCaml "Big'_int.max'_big'_int/ Big'_int.zero'_big'_int") - (Eval "Integer.max/ _/ 0") + (Eval "Integer.max/ 0") text {* For Haskell and Scala, things are slightly different again. *} @@ -339,7 +165,26 @@ (Haskell "toInteger" and "fromInteger") (Scala "!_.as'_BigInt" and "Nat") -text {* Conversion from and to code numerals. *} +text {* Alternativ implementation for @{const of_nat} *} + +lemma [code]: + "of_nat n = (if n = 0 then 0 else + let + (q, m) = divmod_nat n 2; + q' = 2 * of_nat q + in if m = 0 then q' else q' + 1)" +proof - + from mod_div_equality have *: "of_nat n = of_nat (n div 2 * 2 + n mod 2)" by simp + show ?thesis + apply (simp add: Let_def divmod_nat_div_mod mod_2_not_eq_zero_eq_one_nat + of_nat_mult + of_nat_add [symmetric]) + apply (auto simp add: of_nat_mult) + apply (simp add: * of_nat_mult add_commute mult_commute) + done +qed + +text {* Conversion from and to code numerals *} code_const Code_Numeral.of_nat (SML "IntInf.toInt") @@ -355,21 +200,38 @@ (Scala "!Nat(_.as'_BigInt)") (Eval "_") -text {* Using target language arithmetic operations whenever appropriate *} + +subsection {* Target language arithmetic *} -code_const "op + \ nat \ nat \ nat" - (SML "IntInf.+ ((_), (_))") +code_const "plus \ nat \ nat \ nat" + (SML "IntInf.+/ ((_),/ (_))") (OCaml "Big'_int.add'_big'_int") (Haskell infixl 6 "+") (Scala infixl 7 "+") (Eval infixl 8 "+") -code_const "op - \ nat \ nat \ nat" +code_const "minus \ nat \ nat \ nat" + (SML "IntInf.max/ (0, IntInf.-/ ((_),/ (_)))") + (OCaml "Big'_int.max'_big'_int/ Big'_int.zero'_big'_int/ (Big'_int.sub'_big'_int/ _/ _)") (Haskell infixl 6 "-") (Scala infixl 7 "-") + (Eval "Integer.max/ 0/ (_ -/ _)") -code_const "op * \ nat \ nat \ nat" - (SML "IntInf.* ((_), (_))") +code_const Code_Nat.dup + (SML "IntInf.*/ (2,/ (_))") + (OCaml "Big'_int.mult'_big'_int/ 2") + (Haskell "!(2 * _)") + (Scala "!(2 * _)") + (Eval "!(2 * _)") + +code_const Code_Nat.sub + (SML "!(raise/ Fail/ \"sub\")") + (OCaml "failwith/ \"sub\"") + (Haskell "error/ \"sub\"") + (Scala "!error(\"sub\")") + +code_const "times \ nat \ nat \ nat" + (SML "IntInf.*/ ((_),/ (_))") (OCaml "Big'_int.mult'_big'_int") (Haskell infixl 7 "*") (Scala infixl 8 "*") @@ -389,22 +251,28 @@ (Scala infixl 5 "==") (Eval infixl 6 "=") -code_const "op \ \ nat \ nat \ bool" - (SML "IntInf.<= ((_), (_))") +code_const "less_eq \ nat \ nat \ bool" + (SML "IntInf.<=/ ((_),/ (_))") (OCaml "Big'_int.le'_big'_int") (Haskell infix 4 "<=") (Scala infixl 4 "<=") (Eval infixl 6 "<=") -code_const "op < \ nat \ nat \ bool" - (SML "IntInf.< ((_), (_))") +code_const "less \ nat \ nat \ bool" + (SML "IntInf. nat \ term) = Code_Evaluation.term_of" .. @@ -412,14 +280,14 @@ code_const "Code_Evaluation.term_of \ nat \ term" (SML "HOLogic.mk'_number/ HOLogic.natT") -text {* Evaluation with @{text "Quickcheck_Narrowing"} does not work, as +text {* + FIXME -- Evaluation with @{text "Quickcheck_Narrowing"} does not work, as @{text "code_module"} is very aggressive leading to bad Haskell code. Therefore, we simply deactivate the narrowing-based quickcheck from here on. *} declare [[quickcheck_narrowing_active = false]] -text {* Module names *} code_modulename SML Efficient_Nat Arith @@ -430,6 +298,6 @@ code_modulename Haskell Efficient_Nat Arith -hide_const int +hide_const (open) int end diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Extended_Nat.thy --- a/src/HOL/Library/Extended_Nat.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Extended_Nat.thy Mon Mar 26 10:56:56 2012 +0200 @@ -61,19 +61,17 @@ primrec the_enat :: "enat \ nat" where "the_enat (enat n) = n" + subsection {* Constructors and numbers *} -instantiation enat :: "{zero, one, number}" +instantiation enat :: "{zero, one}" begin definition "0 = enat 0" definition - [code_unfold]: "1 = enat 1" - -definition - [code_unfold, code del]: "number_of k = enat (number_of k)" + "1 = enat 1" instance .. @@ -82,15 +80,12 @@ definition eSuc :: "enat \ enat" where "eSuc i = (case i of enat n \ enat (Suc n) | \ \ \)" -lemma enat_0: "enat 0 = 0" +lemma enat_0 [code_post]: "enat 0 = 0" by (simp add: zero_enat_def) -lemma enat_1: "enat 1 = 1" +lemma enat_1 [code_post]: "enat 1 = 1" by (simp add: one_enat_def) -lemma enat_number: "enat (number_of k) = number_of k" - by (simp add: number_of_enat_def) - lemma one_eSuc: "1 = eSuc 0" by (simp add: zero_enat_def one_enat_def eSuc_def) @@ -100,16 +95,6 @@ lemma i0_ne_infinity [simp]: "0 \ (\::enat)" by (simp add: zero_enat_def) -lemma zero_enat_eq [simp]: - "number_of k = (0\enat) \ number_of k = (0\nat)" - "(0\enat) = number_of k \ number_of k = (0\nat)" - unfolding zero_enat_def number_of_enat_def by simp_all - -lemma one_enat_eq [simp]: - "number_of k = (1\enat) \ number_of k = (1\nat)" - "(1\enat) = number_of k \ number_of k = (1\nat)" - unfolding one_enat_def number_of_enat_def by simp_all - lemma zero_one_enat_neq [simp]: "\ 0 = (1\enat)" "\ 1 = (0\enat)" @@ -121,18 +106,9 @@ lemma i1_ne_infinity [simp]: "1 \ (\::enat)" by (simp add: one_enat_def) -lemma infinity_ne_number [simp]: "(\::enat) \ number_of k" - by (simp add: number_of_enat_def) - -lemma number_ne_infinity [simp]: "number_of k \ (\::enat)" - by (simp add: number_of_enat_def) - lemma eSuc_enat: "eSuc (enat n) = enat (Suc n)" by (simp add: eSuc_def) -lemma eSuc_number_of: "eSuc (number_of k) = enat (Suc (number_of k))" - by (simp add: eSuc_enat number_of_enat_def) - lemma eSuc_infinity [simp]: "eSuc \ = \" by (simp add: eSuc_def) @@ -145,11 +121,6 @@ lemma eSuc_inject [simp]: "eSuc m = eSuc n \ m = n" by (simp add: eSuc_def split: enat.splits) -lemma number_of_enat_inject [simp]: - "(number_of k \ enat) = number_of l \ (number_of k \ nat) = number_of l" - by (simp add: number_of_enat_def) - - subsection {* Addition *} instantiation enat :: comm_monoid_add @@ -177,16 +148,6 @@ end -lemma plus_enat_number [simp]: - "(number_of k \ enat) + number_of l = (if k < Int.Pls then number_of l - else if l < Int.Pls then number_of k else number_of (k + l))" - unfolding number_of_enat_def plus_enat_simps nat_arith(1) if_distrib [symmetric, of _ enat] .. - -lemma eSuc_number [simp]: - "eSuc (number_of k) = (if neg (number_of k \ int) then 1 else number_of (Int.succ k))" - unfolding eSuc_number_of - unfolding one_enat_def number_of_enat_def Suc_nat_number_of if_distrib [symmetric] .. - lemma eSuc_plus_1: "eSuc n = n + 1" by (cases n) (simp_all add: eSuc_enat one_enat_def) @@ -261,12 +222,6 @@ apply (simp add: plus_1_eSuc eSuc_enat) done -instance enat :: number_semiring -proof - fix n show "number_of (int n) = (of_nat n :: enat)" - unfolding number_of_enat_def number_of_int of_nat_id of_nat_eq_enat .. -qed - instance enat :: semiring_char_0 proof have "inj enat" by (rule injI) simp then show "inj (\n. of_nat n :: enat)" by (simp add: of_nat_eq_enat) @@ -279,6 +234,25 @@ by (auto simp add: times_enat_def zero_enat_def split: enat.split) +subsection {* Numerals *} + +lemma numeral_eq_enat: + "numeral k = enat (numeral k)" + using of_nat_eq_enat [of "numeral k"] by simp + +lemma enat_numeral [code_abbrev]: + "enat (numeral k) = numeral k" + using numeral_eq_enat .. + +lemma infinity_ne_numeral [simp]: "(\::enat) \ numeral k" + by (simp add: numeral_eq_enat) + +lemma numeral_ne_infinity [simp]: "numeral k \ (\::enat)" + by (simp add: numeral_eq_enat) + +lemma eSuc_numeral [simp]: "eSuc (numeral k) = numeral (k + Num.One)" + by (simp only: eSuc_plus_1 numeral_plus_one) + subsection {* Subtraction *} instantiation enat :: minus @@ -292,13 +266,13 @@ end -lemma idiff_enat_enat [simp,code]: "enat a - enat b = enat (a - b)" +lemma idiff_enat_enat [simp, code]: "enat a - enat b = enat (a - b)" by (simp add: diff_enat_def) -lemma idiff_infinity [simp,code]: "\ - n = (\::enat)" +lemma idiff_infinity [simp, code]: "\ - n = (\::enat)" by (simp add: diff_enat_def) -lemma idiff_infinity_right [simp,code]: "enat a - \ = 0" +lemma idiff_infinity_right [simp, code]: "enat a - \ = 0" by (simp add: diff_enat_def) lemma idiff_0 [simp]: "(0::enat) - n = 0" @@ -344,13 +318,13 @@ "(\::enat) < q \ False" by (simp_all add: less_eq_enat_def less_enat_def split: enat.splits) -lemma number_of_le_enat_iff[simp]: - shows "number_of m \ enat n \ number_of m \ n" -by (auto simp: number_of_enat_def) +lemma numeral_le_enat_iff[simp]: + shows "numeral m \ enat n \ numeral m \ n" +by (auto simp: numeral_eq_enat) -lemma number_of_less_enat_iff[simp]: - shows "number_of m < enat n \ number_of m < n" -by (auto simp: number_of_enat_def) +lemma numeral_less_enat_iff[simp]: + shows "numeral m < enat n \ numeral m < n" +by (auto simp: numeral_eq_enat) lemma enat_ord_code [code]: "enat m \ enat n \ m \ n" @@ -375,10 +349,15 @@ by (simp split: enat.splits) qed +(* BH: These equations are already proven generally for any type in +class linordered_semidom. However, enat is not in that class because +it does not have the cancellation property. Would it be worthwhile to +a generalize linordered_semidom to a new class that includes enat? *) + lemma enat_ord_number [simp]: - "(number_of m \ enat) \ number_of n \ (number_of m \ nat) \ number_of n" - "(number_of m \ enat) < number_of n \ (number_of m \ nat) < number_of n" - by (simp_all add: number_of_enat_def) + "(numeral m \ enat) \ numeral n \ (numeral m \ nat) \ numeral n" + "(numeral m \ enat) < numeral n \ (numeral m \ nat) < numeral n" + by (simp_all add: numeral_eq_enat) lemma i0_lb [simp]: "(0\enat) \ n" by (simp add: zero_enat_def less_eq_enat_def split: enat.splits) @@ -525,10 +504,10 @@ val find_first = find_first_t [] val trans_tac = Numeral_Simprocs.trans_tac val norm_ss = HOL_basic_ss addsimps - @{thms add_ac semiring_numeral_0_eq_0 add_0_left add_0_right} + @{thms add_ac add_0_left add_0_right} fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss)) fun simplify_meta_eq ss cancel_th th = - Arith_Data.simplify_meta_eq @{thms semiring_numeral_0_eq_0} ss + Arith_Data.simplify_meta_eq [] ss ([th, cancel_th] MRS trans) fun mk_eq (a, b) = HOLogic.mk_Trueprop (HOLogic.mk_eq (a, b)) end @@ -646,7 +625,7 @@ subsection {* Traditional theorem names *} -lemmas enat_defs = zero_enat_def one_enat_def number_of_enat_def eSuc_def +lemmas enat_defs = zero_enat_def one_enat_def eSuc_def plus_enat_def less_eq_enat_def less_enat_def end diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Extended_Real.thy --- a/src/HOL/Library/Extended_Real.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Extended_Real.thy Mon Mar 26 10:56:56 2012 +0200 @@ -124,11 +124,6 @@ fix x :: ereal show "x \ range uminus" by (intro image_eqI[of _ _ "-x"]) auto qed auto -instantiation ereal :: number -begin -definition [simp]: "number_of x = ereal (number_of x)" -instance .. -end instantiation ereal :: abs begin @@ -671,6 +666,14 @@ using assms by (cases rule: ereal3_cases[of a b c]) (simp_all add: field_simps) +instance ereal :: numeral .. + +lemma numeral_eq_ereal [simp]: "numeral w = ereal (numeral w)" + apply (induct w rule: num_induct) + apply (simp only: numeral_One one_ereal_def) + apply (simp only: numeral_inc ereal_plus_1) + done + lemma ereal_le_epsilon: fixes x y :: ereal assumes "ALL e. 0 < e --> x <= y + e" @@ -781,8 +784,8 @@ shows "(- x) ^ n = (if even n then x ^ n else - (x^n))" by (induct n) (auto simp: one_ereal_def) -lemma ereal_power_number_of[simp]: - "(number_of num :: ereal) ^ n = ereal (number_of num ^ n)" +lemma ereal_power_numeral[simp]: + "(numeral num :: ereal) ^ n = ereal (numeral num ^ n)" by (induct n) (auto simp: one_ereal_def) lemma zero_le_power_ereal[simp]: @@ -1730,8 +1733,8 @@ "ereal_of_enat m \ ereal_of_enat n \ m \ n" by (cases m n rule: enat2_cases) auto -lemma number_of_le_ereal_of_enat_iff[simp]: - shows "number_of m \ ereal_of_enat n \ number_of m \ n" +lemma numeral_le_ereal_of_enat_iff[simp]: + shows "numeral m \ ereal_of_enat n \ numeral m \ n" by (cases n) (auto dest: natceiling_le intro: natceiling_le_eq[THEN iffD1]) lemma ereal_of_enat_ge_zero_cancel_iff[simp]: diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Float.thy --- a/src/HOL/Library/Float.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Float.thy Mon Mar 26 10:56:56 2012 +0200 @@ -41,18 +41,6 @@ instance .. end -instantiation float :: number -begin -definition number_of_float where "number_of n = Float n 0" -instance .. -end - -lemma number_of_float_Float: - "number_of k = Float (number_of k) 0" - by (simp add: number_of_float_def number_of_is_id) - -declare number_of_float_Float [symmetric, code_abbrev] - lemma real_of_float_simp[simp]: "real (Float a b) = real a * pow2 b" unfolding real_of_float_def using of_float.simps . @@ -63,12 +51,9 @@ lemma Float_num[simp]: shows "real (Float 1 0) = 1" and "real (Float 1 1) = 2" and "real (Float 1 2) = 4" and "real (Float 1 -1) = 1/2" and "real (Float 1 -2) = 1/4" and "real (Float 1 -3) = 1/8" and - "real (Float -1 0) = -1" and "real (Float (number_of n) 0) = number_of n" + "real (Float -1 0) = -1" and "real (Float (numeral n) 0) = numeral n" by auto -lemma float_number_of[simp]: "real (number_of x :: float) = number_of x" - by (simp only:number_of_float_def Float_num[unfolded number_of_is_id]) - lemma float_number_of_int[simp]: "real (Float n 0) = real n" by simp @@ -349,6 +334,21 @@ by (cases a, cases b) (simp add: plus_float.simps) qed +instance float :: numeral .. + +lemma Float_add_same_scale: "Float x e + Float y e = Float (x + y) e" + by (simp add: plus_float.simps) + +(* FIXME: define other constant for code_unfold_post *) +lemma numeral_float_Float (*[code_unfold_post]*): + "numeral k = Float (numeral k) 0" + by (induct k, simp_all only: numeral.simps one_float_def + Float_add_same_scale) + +lemma float_number_of[simp]: "real (numeral x :: float) = numeral x" + by (simp only: numeral_float_Float Float_num) + + instance float :: comm_monoid_mult proof (intro_classes) fix a b c :: float @@ -555,6 +555,7 @@ show ?thesis unfolding real_of_float_nge0_exp[OF P] divide_inverse by auto qed +(* BROKEN lemma bitlen_Pls: "bitlen (Int.Pls) = Int.Pls" by (subst Pls_def, subst Pls_def, simp) lemma bitlen_Min: "bitlen (Int.Min) = Int.Bit1 Int.Pls" by (subst Min_def, simp add: Bit1_def) @@ -588,6 +589,7 @@ lemma bitlen_number_of: "bitlen (number_of w) = number_of (bitlen w)" by (simp add: number_of_is_id) +BH *) lemma [code]: "bitlen x = (if x = 0 then 0 @@ -722,12 +724,12 @@ hence "real x / real y < 1" using `0 < y` and `0 \ x` by auto from real_of_int_div4[of "?X" y] - have "real (?X div y) \ (real x / real y) * 2^?l" unfolding real_of_int_mult times_divide_eq_left real_of_int_power real_number_of . + have "real (?X div y) \ (real x / real y) * 2^?l" unfolding real_of_int_mult times_divide_eq_left real_of_int_power real_numeral . also have "\ < 1 * 2^?l" using `real x / real y < 1` by (rule mult_strict_right_mono, auto) finally have "?X div y < 2^?l" unfolding real_of_int_less_iff[of _ "2^?l", symmetric] by auto hence "?X div y + 1 \ 2^?l" by auto hence "real (?X div y + 1) * inverse (2^?l) \ 2^?l * inverse (2^?l)" - unfolding real_of_int_le_iff[of _ "2^?l", symmetric] real_of_int_power real_number_of + unfolding real_of_int_le_iff[of _ "2^?l", symmetric] real_of_int_power real_numeral by (rule mult_right_mono, auto) hence "real (?X div y + 1) * inverse (2^?l) \ 1" by auto thus ?thesis unfolding rapprox_posrat_def Let_def normfloat real_of_float_simp if_not_P[OF False] @@ -796,12 +798,12 @@ qed from real_of_int_div4[of "?X" y] - have "real (?X div y) \ (real x / real y) * 2^?l" unfolding real_of_int_mult times_divide_eq_left real_of_int_power real_number_of . + have "real (?X div y) \ (real x / real y) * 2^?l" unfolding real_of_int_mult times_divide_eq_left real_of_int_power real_numeral . also have "\ < 1/2 * 2^?l" using `real x / real y < 1/2` by (rule mult_strict_right_mono, auto) finally have "?X div y * 2 < 2^?l" unfolding real_of_int_less_iff[of _ "2^?l", symmetric] by auto hence "?X div y + 1 < 2^?l" using `0 < ?X div y` by auto hence "real (?X div y + 1) * inverse (2^?l) < 2^?l * inverse (2^?l)" - unfolding real_of_int_less_iff[of _ "2^?l", symmetric] real_of_int_power real_number_of + unfolding real_of_int_less_iff[of _ "2^?l", symmetric] real_of_int_power real_numeral by (rule mult_strict_right_mono, auto) hence "real (?X div y + 1) * inverse (2^?l) < 1" by auto thus ?thesis unfolding rapprox_posrat_def Let_def normfloat real_of_float_simp if_not_P[OF False] @@ -1195,7 +1197,7 @@ case True have "real (m div 2^(nat ?l)) * pow2 ?l \ real m" proof - - have "real (m div 2^(nat ?l)) * pow2 ?l = real (2^(nat ?l) * (m div 2^(nat ?l)))" unfolding real_of_int_mult real_of_int_power real_number_of unfolding pow2_int[symmetric] + have "real (m div 2^(nat ?l)) * pow2 ?l = real (2^(nat ?l) * (m div 2^(nat ?l)))" unfolding real_of_int_mult real_of_int_power real_numeral unfolding pow2_int[symmetric] using `?l > 0` by auto also have "\ \ real (2^(nat ?l) * (m div 2^(nat ?l)) + m mod 2^(nat ?l))" unfolding real_of_int_add by auto also have "\ = real m" unfolding zmod_zdiv_equality[symmetric] .. @@ -1262,7 +1264,7 @@ hence me_eq: "pow2 (-e) = pow2 (int (nat (-e)))" by auto have "real (Float (m div (2 ^ (nat (-e)))) 0) = real (m div 2 ^ (nat (-e)))" unfolding real_of_float_simp by auto also have "\ \ real m / real ((2::int) ^ (nat (-e)))" using real_of_int_div4 . - also have "\ = real m * inverse (2 ^ (nat (-e)))" unfolding real_of_int_power real_number_of divide_inverse .. + also have "\ = real m * inverse (2 ^ (nat (-e)))" unfolding real_of_int_power real_numeral divide_inverse .. also have "\ = real (Float m e)" unfolding real_of_float_simp me_eq pow2_int pow2_neg[of e] .. finally show ?thesis unfolding Float floor_fl.simps if_not_P[OF `\ 0 \ e`] . next @@ -1290,7 +1292,7 @@ case False hence me_eq: "pow2 (-e) = pow2 (int (nat (-e)))" by auto have "real (Float m e) = real m * inverse (2 ^ (nat (-e)))" unfolding real_of_float_simp me_eq pow2_int pow2_neg[of e] .. - also have "\ = real m / real ((2::int) ^ (nat (-e)))" unfolding real_of_int_power real_number_of divide_inverse .. + also have "\ = real m / real ((2::int) ^ (nat (-e)))" unfolding real_of_int_power real_numeral divide_inverse .. also have "\ \ 1 + real (m div 2 ^ (nat (-e)))" using real_of_int_div3[unfolded diff_le_eq] . also have "\ = real (Float (m div (2 ^ (nat (-e))) + 1) 0)" unfolding real_of_float_simp by auto finally show ?thesis unfolding Float ceiling_fl.simps if_not_P[OF `\ 0 \ e`] . diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Formal_Power_Series.thy --- a/src/HOL/Library/Formal_Power_Series.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Formal_Power_Series.thy Mon Mar 26 10:56:56 2012 +0200 @@ -392,25 +392,13 @@ instance fps :: (idom) idom .. -instantiation fps :: (comm_ring_1) number_ring -begin -definition number_of_fps_def: "(number_of k::'a fps) = of_int k" - -instance proof -qed (rule number_of_fps_def) -end - -lemma number_of_fps_const: "(number_of k::('a::comm_ring_1) fps) = fps_const (of_int k)" - -proof(induct k rule: int_induct [where k=0]) - case base thus ?case unfolding number_of_fps_def of_int_0 by simp -next - case (step1 i) thus ?case unfolding number_of_fps_def - by (simp add: fps_const_add[symmetric] del: fps_const_add) -next - case (step2 i) thus ?case unfolding number_of_fps_def - by (simp add: fps_const_sub[symmetric] del: fps_const_sub) -qed +lemma numeral_fps_const: "numeral k = fps_const (numeral k)" + by (induct k, simp_all only: numeral.simps fps_const_1_eq_1 + fps_const_add [symmetric]) + +lemma neg_numeral_fps_const: "neg_numeral k = fps_const (neg_numeral k)" + by (simp only: neg_numeral_def numeral_fps_const fps_const_neg) + subsection{* The eXtractor series X*} lemma minus_one_power_iff: "(- (1::'a :: {comm_ring_1})) ^ n = (if even n then 1 else - 1)" @@ -1119,7 +1107,7 @@ have eq: "(1 + X) * ?r = 1" unfolding minus_one_power_iff by (auto simp add: field_simps fps_eq_iff) - show ?thesis by (auto simp add: eq intro: fps_inverse_unique) + show ?thesis by (auto simp add: eq intro: fps_inverse_unique simp del: minus_one) qed @@ -1157,8 +1145,11 @@ "fps_const (a::'a::{comm_ring_1}) oo b = fps_const (a)" by (simp add: fps_eq_iff fps_compose_nth mult_delta_left setsum_delta) -lemma number_of_compose[simp]: "(number_of k::('a::{comm_ring_1}) fps) oo b = number_of k" - unfolding number_of_fps_const by simp +lemma numeral_compose[simp]: "(numeral k::('a::{comm_ring_1}) fps) oo b = numeral k" + unfolding numeral_fps_const by simp + +lemma neg_numeral_compose[simp]: "(neg_numeral k::('a::{comm_ring_1}) fps) oo b = neg_numeral k" + unfolding neg_numeral_fps_const by simp lemma X_fps_compose_startby0[simp]: "a$0 = 0 \ X oo a = (a :: ('a :: comm_ring_1) fps)" by (simp add: fps_eq_iff fps_compose_def mult_delta_left setsum_delta @@ -2568,7 +2559,7 @@ (is "inverse ?l = ?r") proof- have th: "?l * ?r = 1" - by (auto simp add: field_simps fps_eq_iff minus_one_power_iff) + by (auto simp add: field_simps fps_eq_iff minus_one_power_iff simp del: minus_one) have th': "?l $ 0 \ 0" by (simp add: ) from fps_inverse_unique[OF th' th] show ?thesis . qed @@ -2765,7 +2756,7 @@ proof- have th: "?r$0 \ 0" by simp have th': "fps_deriv (inverse ?r) = fps_const (- 1) * inverse ?r / (1 + X)" - by (simp add: fps_inverse_deriv[OF th] fps_divide_def power2_eq_square mult_commute fps_const_neg[symmetric] del: fps_const_neg) + by (simp add: fps_inverse_deriv[OF th] fps_divide_def power2_eq_square mult_commute fps_const_neg[symmetric] del: fps_const_neg minus_one) have eq: "inverse ?r $ 0 = 1" by (simp add: fps_inverse_def) from iffD1[OF fps_binomial_ODE_unique[of "inverse (1 + X)" "- 1"] th'] eq @@ -2855,7 +2846,7 @@ unfolding m1nk unfolding m h pochhammer_Suc_setprod - apply (simp add: field_simps del: fact_Suc id_def) + apply (simp add: field_simps del: fact_Suc id_def minus_one) unfolding fact_altdef_nat id_def unfolding of_nat_setprod unfolding setprod_timesf[symmetric] @@ -3162,28 +3153,25 @@ lemma fps_const_minus: "fps_const (c::'a::group_add) - fps_const d = fps_const (c - d)" by (simp add: fps_eq_iff fps_const_def) -lemma fps_number_of_fps_const: "number_of i = fps_const (number_of i :: 'a:: {comm_ring_1, number_ring})" - apply (subst (2) number_of_eq) -apply(rule int_induct [of _ 0]) -apply (simp_all add: number_of_fps_def) -by (simp_all add: fps_const_add[symmetric] fps_const_minus[symmetric]) +lemma fps_numeral_fps_const: "numeral i = fps_const (numeral i :: 'a:: {comm_ring_1})" + by (fact numeral_fps_const) (* FIXME: duplicate *) lemma fps_cos_Eii: "fps_cos c = (E (ii * c) + E (- ii * c)) / fps_const 2" proof- have th: "fps_cos c + fps_cos c = fps_cos c * fps_const 2" - by (simp add: fps_eq_iff fps_number_of_fps_const complex_number_of_def[symmetric]) + by (simp add: numeral_fps_const) show ?thesis unfolding Eii_sin_cos minus_mult_commute - by (simp add: fps_sin_even fps_cos_odd fps_number_of_fps_const - fps_divide_def fps_const_inverse th complex_number_of_def[symmetric]) + by (simp add: fps_sin_even fps_cos_odd numeral_fps_const + fps_divide_def fps_const_inverse th) qed lemma fps_sin_Eii: "fps_sin c = (E (ii * c) - E (- ii * c)) / fps_const (2*ii)" proof- have th: "fps_const \ * fps_sin c + fps_const \ * fps_sin c = fps_sin c * fps_const (2 * ii)" - by (simp add: fps_eq_iff fps_number_of_fps_const complex_number_of_def[symmetric]) + by (simp add: fps_eq_iff numeral_fps_const) show ?thesis unfolding Eii_sin_cos minus_mult_commute by (simp add: fps_sin_even fps_cos_odd fps_divide_def fps_const_inverse th) diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Numeral_Type.thy --- a/src/HOL/Library/Numeral_Type.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Numeral_Type.thy Mon Mar 26 10:56:56 2012 +0200 @@ -66,7 +66,6 @@ by simp qed - subsection {* Locales for for modular arithmetic subtypes *} locale mod_type = @@ -137,8 +136,8 @@ locale mod_ring = mod_type n Rep Abs for n :: int - and Rep :: "'a::{number_ring} \ int" - and Abs :: "int \ 'a::{number_ring}" + and Rep :: "'a::{comm_ring_1} \ int" + and Abs :: "int \ 'a::{comm_ring_1}" begin lemma of_nat_eq: "of_nat k = Abs (int k mod n)" @@ -152,13 +151,14 @@ apply (simp add: Rep_simps of_nat_eq diff_def zmod_simps) done -lemma Rep_number_of: - "Rep (number_of w) = number_of w mod n" -by (simp add: number_of_eq of_int_eq Rep_Abs_mod) +lemma Rep_numeral: + "Rep (numeral w) = numeral w mod n" +using of_int_eq [of "numeral w"] +by (simp add: Rep_inject_sym Rep_Abs_mod) -lemma iszero_number_of: - "iszero (number_of w::'a) \ number_of w mod n = 0" -by (simp add: Rep_simps number_of_eq of_int_eq iszero_def zero_def) +lemma iszero_numeral: + "iszero (numeral w::'a) \ numeral w mod n = 0" +by (simp add: Rep_inject_sym Rep_numeral Rep_0 iszero_def) lemma cases: assumes 1: "\z. \(x::'a) = of_int z; 0 \ z; z < n\ \ P" @@ -175,14 +175,14 @@ end -subsection {* Number ring instances *} +subsection {* Ring class instances *} text {* - Unfortunately a number ring instance is not possible for + Unfortunately @{text ring_1} instance is not possible for @{typ num1}, since 0 and 1 are not distinct. *} -instantiation num1 :: "{comm_ring,comm_monoid_mult,number}" +instantiation num1 :: "{comm_ring,comm_monoid_mult,numeral}" begin lemma num1_eq_iff: "(x::num1) = (y::num1) \ True" @@ -252,22 +252,10 @@ done instance bit0 :: (finite) comm_ring_1 - by (rule bit0.comm_ring_1)+ + by (rule bit0.comm_ring_1) instance bit1 :: (finite) comm_ring_1 - by (rule bit1.comm_ring_1)+ - -instantiation bit0 and bit1 :: (finite) number_ring -begin - -definition "(number_of w :: _ bit0) = of_int w" - -definition "(number_of w :: _ bit1) = of_int w" - -instance proof -qed (rule number_of_bit0_def number_of_bit1_def)+ - -end + by (rule bit1.comm_ring_1) interpretation bit0: mod_ring "int CARD('a::finite bit0)" @@ -289,9 +277,11 @@ lemmas bit0_induct [case_names of_int, induct type: bit0] = bit0.induct lemmas bit1_induct [case_names of_int, induct type: bit1] = bit1.induct -lemmas bit0_iszero_number_of [simp] = bit0.iszero_number_of -lemmas bit1_iszero_number_of [simp] = bit1.iszero_number_of +lemmas bit0_iszero_numeral [simp] = bit0.iszero_numeral +lemmas bit1_iszero_numeral [simp] = bit1.iszero_numeral +declare eq_numeral_iff_iszero [where 'a="('a::finite) bit0", standard, simp] +declare eq_numeral_iff_iszero [where 'a="('a::finite) bit1", standard, simp] subsection {* Syntax *} diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Poly_Deriv.thy --- a/src/HOL/Library/Poly_Deriv.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Poly_Deriv.thy Mon Mar 26 10:56:56 2012 +0200 @@ -71,7 +71,8 @@ apply (subst power_Suc) apply (subst pderiv_mult) apply (erule ssubst) -apply (simp add: smult_add_left algebra_simps) +apply (simp only: of_nat_Suc smult_add_left smult_1_left) +apply (simp add: algebra_simps) (* FIXME *) done lemma DERIV_cmult2: "DERIV f x :> D ==> DERIV (%x. (f x) * c :: real) x :> D * c" diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Polynomial.thy --- a/src/HOL/Library/Polynomial.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Polynomial.thy Mon Mar 26 10:56:56 2012 +0200 @@ -662,17 +662,6 @@ instance poly :: (comm_ring_1) comm_ring_1 .. -instantiation poly :: (comm_ring_1) number_ring -begin - -definition - "number_of k = (of_int k :: 'a poly)" - -instance - by default (rule number_of_poly_def) - -end - subsection {* Polynomials form an integral domain *} @@ -1052,12 +1041,12 @@ lemma poly_div_minus_left [simp]: fixes x y :: "'a::field poly" shows "(- x) div y = - (x div y)" - using div_smult_left [of "- 1::'a"] by simp + using div_smult_left [of "- 1::'a"] by (simp del: minus_one) (* FIXME *) lemma poly_mod_minus_left [simp]: fixes x y :: "'a::field poly" shows "(- x) mod y = - (x mod y)" - using mod_smult_left [of "- 1::'a"] by simp + using mod_smult_left [of "- 1::'a"] by (simp del: minus_one) (* FIXME *) lemma pdivmod_rel_smult_right: "\a \ 0; pdivmod_rel x y q r\ @@ -1075,12 +1064,12 @@ fixes x y :: "'a::field poly" shows "x div (- y) = - (x div y)" using div_smult_right [of "- 1::'a"] - by (simp add: nonzero_inverse_minus_eq) + by (simp add: nonzero_inverse_minus_eq del: minus_one) (* FIXME *) lemma poly_mod_minus_right [simp]: fixes x y :: "'a::field poly" shows "x mod (- y) = x mod y" - using mod_smult_right [of "- 1::'a"] by simp + using mod_smult_right [of "- 1::'a"] by (simp del: minus_one) (* FIXME *) lemma pdivmod_rel_mult: "\pdivmod_rel x y q r; pdivmod_rel q z q' r'\ diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Predicate_Compile_Alternative_Defs.thy --- a/src/HOL/Library/Predicate_Compile_Alternative_Defs.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Predicate_Compile_Alternative_Defs.thy Mon Mar 26 10:56:56 2012 +0200 @@ -54,8 +54,8 @@ section {* Setup for Numerals *} -setup {* Predicate_Compile_Data.ignore_consts [@{const_name number_of}] *} -setup {* Predicate_Compile_Data.keep_functions [@{const_name number_of}] *} +setup {* Predicate_Compile_Data.ignore_consts [@{const_name numeral}, @{const_name neg_numeral}] *} +setup {* Predicate_Compile_Data.keep_functions [@{const_name numeral}, @{const_name neg_numeral}] *} setup {* Predicate_Compile_Data.ignore_consts [@{const_name div}, @{const_name mod}, @{const_name times}] *} diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Quotient_List.thy --- a/src/HOL/Library/Quotient_List.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Quotient_List.thy Mon Mar 26 10:56:56 2012 +0200 @@ -8,8 +8,6 @@ imports Main Quotient_Syntax begin -declare [[map list = list_all2]] - lemma map_id [id_simps]: "map id = id" by (fact List.map.id) @@ -75,6 +73,8 @@ by (induct xs ys rule: list_induct2') auto qed +declare [[map list = (list_all2, list_quotient)]] + lemma cons_prs [quot_preserve]: assumes q: "Quotient R Abs Rep" shows "(Rep ---> (map Rep) ---> (map Abs)) (op #) = (op #)" diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Quotient_Option.thy --- a/src/HOL/Library/Quotient_Option.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Quotient_Option.thy Mon Mar 26 10:56:56 2012 +0200 @@ -16,8 +16,6 @@ | "option_rel R None (Some x) = False" | "option_rel R (Some x) (Some y) = R x y" -declare [[map option = option_rel]] - lemma option_rel_unfold: "option_rel R x y = (case (x, y) of (None, None) \ True | (Some x, Some y) \ R x y @@ -65,6 +63,8 @@ apply (simp add: option_rel_unfold split: option.split) done +declare [[map option = (option_rel, option_quotient)]] + lemma option_None_rsp [quot_respect]: assumes q: "Quotient R Abs Rep" shows "option_rel R None None" diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Quotient_Product.thy --- a/src/HOL/Library/Quotient_Product.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Quotient_Product.thy Mon Mar 26 10:56:56 2012 +0200 @@ -13,8 +13,6 @@ where "prod_rel R1 R2 = (\(a, b) (c, d). R1 a c \ R2 b d)" -declare [[map prod = prod_rel]] - lemma prod_rel_apply [simp]: "prod_rel R1 R2 (a, b) (c, d) \ R1 a c \ R2 b d" by (simp add: prod_rel_def) @@ -45,6 +43,8 @@ apply (auto simp add: split_paired_all) done +declare [[map prod = (prod_rel, prod_quotient)]] + lemma Pair_rsp [quot_respect]: assumes q1: "Quotient R1 Abs1 Rep1" assumes q2: "Quotient R2 Abs2 Rep2" diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Quotient_Set.thy --- a/src/HOL/Library/Quotient_Set.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Quotient_Set.thy Mon Mar 26 10:56:56 2012 +0200 @@ -26,6 +26,8 @@ by auto (metis rep_abs_rsp[OF assms] assms[simplified Quotient_def])+ qed +declare [[map set = (set_rel, set_quotient)]] + lemma empty_set_rsp[quot_respect]: "set_rel R {} {}" unfolding set_rel_def by simp diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Quotient_Sum.thy --- a/src/HOL/Library/Quotient_Sum.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Quotient_Sum.thy Mon Mar 26 10:56:56 2012 +0200 @@ -16,8 +16,6 @@ | "sum_rel R1 R2 (Inr a2) (Inl b1) = False" | "sum_rel R1 R2 (Inr a2) (Inr b2) = R2 a2 b2" -declare [[map sum = sum_rel]] - lemma sum_rel_unfold: "sum_rel R1 R2 x y = (case (x, y) of (Inl x, Inl y) \ R1 x y | (Inr x, Inr y) \ R2 x y @@ -67,6 +65,8 @@ apply (simp add: sum_rel_unfold comp_def split: sum.split) done +declare [[map sum = (sum_rel, sum_quotient)]] + lemma sum_Inl_rsp [quot_respect]: assumes q1: "Quotient R1 Abs1 Rep1" assumes q2: "Quotient R2 Abs2 Rep2" diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/ROOT.ML --- a/src/HOL/Library/ROOT.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/ROOT.ML Mon Mar 26 10:56:56 2012 +0200 @@ -4,4 +4,4 @@ use_thys ["Library", "List_Cset", "List_Prefix", "List_lexord", "Sublist_Order", "Product_Lattice", "Code_Char_chr", "Code_Char_ord", "Code_Integer", "Efficient_Nat"(*, "Code_Prolog"*), - "Code_Real_Approx_By_Float" ]; + "Code_Real_Approx_By_Float", "Target_Numeral"]; diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Saturated.thy --- a/src/HOL/Library/Saturated.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Saturated.thy Mon Mar 26 10:56:56 2012 +0200 @@ -157,20 +157,16 @@ "nat_of (Sat n :: ('a::len) sat) = min (len_of TYPE('a)) n" by (rule nat_of_Abs_sat' [unfolded Abs_sat'_eq_of_nat]) -instantiation sat :: (len) number_semiring -begin +lemma [code_abbrev]: + "of_nat (numeral k) = (numeral k :: 'a::len sat)" + by simp -definition - number_of_sat_def [code del]: "number_of = Sat \ nat" - -instance - by default (simp add: number_of_sat_def) - -end +definition sat_of_nat :: "nat \ ('a::len) sat" + where [code_abbrev]: "sat_of_nat = of_nat" lemma [code abstract]: - "nat_of (number_of n :: ('a::len) sat) = min (nat n) (len_of TYPE('a))" - unfolding number_of_sat_def by simp + "nat_of (sat_of_nat n :: ('a::len) sat) = min (len_of TYPE('a)) n" + by (simp add: sat_of_nat_def) instance sat :: (len) finite proof @@ -252,4 +248,6 @@ end +hide_const (open) sat_of_nat + end diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Sum_of_Squares/sum_of_squares.ML --- a/src/HOL/Library/Sum_of_Squares/sum_of_squares.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Library/Sum_of_Squares/sum_of_squares.ML Mon Mar 26 10:56:56 2012 +0200 @@ -866,10 +866,11 @@ @{term "op / :: real => _"}, @{term "inverse :: real => _"}, @{term "op ^ :: real => _"}, @{term "abs :: real => _"}, @{term "min :: real => _"}, @{term "max :: real => _"}, - @{term "0::real"}, @{term "1::real"}, @{term "number_of :: int => real"}, - @{term "number_of :: int => nat"}, - @{term "Int.Bit0"}, @{term "Int.Bit1"}, - @{term "Int.Pls"}, @{term "Int.Min"}]; + @{term "0::real"}, @{term "1::real"}, + @{term "numeral :: num => nat"}, + @{term "numeral :: num => real"}, + @{term "neg_numeral :: num => real"}, + @{term "Num.Bit0"}, @{term "Num.Bit1"}, @{term "Num.One"}]; fun check_sos kcts ct = let diff -r 29e92b644d6c -r f067afe98049 src/HOL/Library/Target_Numeral.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Target_Numeral.thy Mon Mar 26 10:56:56 2012 +0200 @@ -0,0 +1,726 @@ +theory Target_Numeral +imports Main Code_Nat +begin + +subsection {* Type of target language numerals *} + +typedef (open) int = "UNIV \ int set" + morphisms int_of of_int .. + +hide_type (open) int +hide_const (open) of_int + +lemma int_eq_iff: + "k = l \ int_of k = int_of l" + using int_of_inject [of k l] .. + +lemma int_eqI: + "int_of k = int_of l \ k = l" + using int_eq_iff [of k l] by simp + +lemma int_of_int [simp]: + "int_of (Target_Numeral.of_int k) = k" + using of_int_inverse [of k] by simp + +lemma of_int_of [simp]: + "Target_Numeral.of_int (int_of k) = k" + using int_of_inverse [of k] by simp + +hide_fact (open) int_eq_iff int_eqI + +instantiation Target_Numeral.int :: ring_1 +begin + +definition + "0 = Target_Numeral.of_int 0" + +lemma int_of_zero [simp]: + "int_of 0 = 0" + by (simp add: zero_int_def) + +definition + "1 = Target_Numeral.of_int 1" + +lemma int_of_one [simp]: + "int_of 1 = 1" + by (simp add: one_int_def) + +definition + "k + l = Target_Numeral.of_int (int_of k + int_of l)" + +lemma int_of_plus [simp]: + "int_of (k + l) = int_of k + int_of l" + by (simp add: plus_int_def) + +definition + "- k = Target_Numeral.of_int (- int_of k)" + +lemma int_of_uminus [simp]: + "int_of (- k) = - int_of k" + by (simp add: uminus_int_def) + +definition + "k - l = Target_Numeral.of_int (int_of k - int_of l)" + +lemma int_of_minus [simp]: + "int_of (k - l) = int_of k - int_of l" + by (simp add: minus_int_def) + +definition + "k * l = Target_Numeral.of_int (int_of k * int_of l)" + +lemma int_of_times [simp]: + "int_of (k * l) = int_of k * int_of l" + by (simp add: times_int_def) + +instance proof +qed (auto simp add: Target_Numeral.int_eq_iff algebra_simps) + +end + +lemma int_of_of_nat [simp]: + "int_of (of_nat n) = of_nat n" + by (induct n) simp_all + +definition nat_of :: "Target_Numeral.int \ nat" where + "nat_of k = Int.nat (int_of k)" + +lemma nat_of_of_nat [simp]: + "nat_of (of_nat n) = n" + by (simp add: nat_of_def) + +lemma int_of_of_int [simp]: + "int_of (of_int k) = k" + by (induct k) (simp_all, simp only: neg_numeral_def numeral_One int_of_uminus int_of_one) + +lemma of_int_of_int [simp, code_abbrev]: + "Target_Numeral.of_int = of_int" + by rule (simp add: Target_Numeral.int_eq_iff) + +lemma int_of_numeral [simp]: + "int_of (numeral k) = numeral k" + using int_of_of_int [of "numeral k"] by simp + +lemma int_of_neg_numeral [simp]: + "int_of (neg_numeral k) = neg_numeral k" + by (simp only: neg_numeral_def int_of_uminus) simp + +lemma int_of_sub [simp]: + "int_of (Num.sub k l) = Num.sub k l" + by (simp only: Num.sub_def int_of_minus int_of_numeral) + +instantiation Target_Numeral.int :: "{ring_div, equal, linordered_idom}" +begin + +definition + "k div l = of_int (int_of k div int_of l)" + +lemma int_of_div [simp]: + "int_of (k div l) = int_of k div int_of l" + by (simp add: div_int_def) + +definition + "k mod l = of_int (int_of k mod int_of l)" + +lemma int_of_mod [simp]: + "int_of (k mod l) = int_of k mod int_of l" + by (simp add: mod_int_def) + +definition + "\k\ = of_int \int_of k\" + +lemma int_of_abs [simp]: + "int_of \k\ = \int_of k\" + by (simp add: abs_int_def) + +definition + "sgn k = of_int (sgn (int_of k))" + +lemma int_of_sgn [simp]: + "int_of (sgn k) = sgn (int_of k)" + by (simp add: sgn_int_def) + +definition + "k \ l \ int_of k \ int_of l" + +definition + "k < l \ int_of k < int_of l" + +definition + "HOL.equal k l \ HOL.equal (int_of k) (int_of l)" + +instance proof +qed (auto simp add: Target_Numeral.int_eq_iff algebra_simps + less_eq_int_def less_int_def equal_int_def equal) + +end + +lemma int_of_min [simp]: + "int_of (min k l) = min (int_of k) (int_of l)" + by (simp add: min_def less_eq_int_def) + +lemma int_of_max [simp]: + "int_of (max k l) = max (int_of k) (int_of l)" + by (simp add: max_def less_eq_int_def) + + +subsection {* Code theorems for target language numerals *} + +text {* Constructors *} + +definition Pos :: "num \ Target_Numeral.int" where + [simp, code_abbrev]: "Pos = numeral" + +definition Neg :: "num \ Target_Numeral.int" where + [simp, code_abbrev]: "Neg = neg_numeral" + +code_datatype "0::Target_Numeral.int" Pos Neg + + +text {* Auxiliary operations *} + +definition dup :: "Target_Numeral.int \ Target_Numeral.int" where + [simp]: "dup k = k + k" + +lemma dup_code [code]: + "dup 0 = 0" + "dup (Pos n) = Pos (Num.Bit0 n)" + "dup (Neg n) = Neg (Num.Bit0 n)" + unfolding Pos_def Neg_def neg_numeral_def + by (simp_all add: numeral_Bit0) + +definition sub :: "num \ num \ Target_Numeral.int" where + [simp]: "sub m n = numeral m - numeral n" + +lemma sub_code [code]: + "sub Num.One Num.One = 0" + "sub (Num.Bit0 m) Num.One = Pos (Num.BitM m)" + "sub (Num.Bit1 m) Num.One = Pos (Num.Bit0 m)" + "sub Num.One (Num.Bit0 n) = Neg (Num.BitM n)" + "sub Num.One (Num.Bit1 n) = Neg (Num.Bit0 n)" + "sub (Num.Bit0 m) (Num.Bit0 n) = dup (sub m n)" + "sub (Num.Bit1 m) (Num.Bit1 n) = dup (sub m n)" + "sub (Num.Bit1 m) (Num.Bit0 n) = dup (sub m n) + 1" + "sub (Num.Bit0 m) (Num.Bit1 n) = dup (sub m n) - 1" + unfolding sub_def dup_def numeral.simps Pos_def Neg_def + neg_numeral_def numeral_BitM + by (simp_all only: algebra_simps add.comm_neutral) + + +text {* Implementations *} + +lemma one_int_code [code, code_unfold]: + "1 = Pos Num.One" + by simp + +lemma plus_int_code [code]: + "k + 0 = (k::Target_Numeral.int)" + "0 + l = (l::Target_Numeral.int)" + "Pos m + Pos n = Pos (m + n)" + "Pos m + Neg n = sub m n" + "Neg m + Pos n = sub n m" + "Neg m + Neg n = Neg (m + n)" + by simp_all + +lemma uminus_int_code [code]: + "uminus 0 = (0::Target_Numeral.int)" + "uminus (Pos m) = Neg m" + "uminus (Neg m) = Pos m" + by simp_all + +lemma minus_int_code [code]: + "k - 0 = (k::Target_Numeral.int)" + "0 - l = uminus (l::Target_Numeral.int)" + "Pos m - Pos n = sub m n" + "Pos m - Neg n = Pos (m + n)" + "Neg m - Pos n = Neg (m + n)" + "Neg m - Neg n = sub n m" + by simp_all + +lemma times_int_code [code]: + "k * 0 = (0::Target_Numeral.int)" + "0 * l = (0::Target_Numeral.int)" + "Pos m * Pos n = Pos (m * n)" + "Pos m * Neg n = Neg (m * n)" + "Neg m * Pos n = Neg (m * n)" + "Neg m * Neg n = Pos (m * n)" + by simp_all + +definition divmod :: "Target_Numeral.int \ Target_Numeral.int \ Target_Numeral.int \ Target_Numeral.int" where + "divmod k l = (k div l, k mod l)" + +lemma fst_divmod [simp]: + "fst (divmod k l) = k div l" + by (simp add: divmod_def) + +lemma snd_divmod [simp]: + "snd (divmod k l) = k mod l" + by (simp add: divmod_def) + +definition divmod_abs :: "Target_Numeral.int \ Target_Numeral.int \ Target_Numeral.int \ Target_Numeral.int" where + "divmod_abs k l = (\k\ div \l\, \k\ mod \l\)" + +lemma fst_divmod_abs [simp]: + "fst (divmod_abs k l) = \k\ div \l\" + by (simp add: divmod_abs_def) + +lemma snd_divmod_abs [simp]: + "snd (divmod_abs k l) = \k\ mod \l\" + by (simp add: divmod_abs_def) + +lemma divmod_abs_terminate_code [code]: + "divmod_abs (Neg k) (Neg l) = divmod_abs (Pos k) (Pos l)" + "divmod_abs (Neg k) (Pos l) = divmod_abs (Pos k) (Pos l)" + "divmod_abs (Pos k) (Neg l) = divmod_abs (Pos k) (Pos l)" + "divmod_abs j 0 = (0, \j\)" + "divmod_abs 0 j = (0, 0)" + by (simp_all add: prod_eq_iff) + +lemma divmod_abs_rec_code [code]: + "divmod_abs (Pos k) (Pos l) = + (let j = sub k l in + if j < 0 then (0, Pos k) + else let (q, r) = divmod_abs j (Pos l) in (q + 1, r))" + by (auto simp add: prod_eq_iff Target_Numeral.int_eq_iff Let_def prod_case_beta + sub_non_negative sub_negative div_pos_pos_trivial mod_pos_pos_trivial div_pos_geq mod_pos_geq) + +lemma divmod_code [code]: "divmod k l = + (if k = 0 then (0, 0) else if l = 0 then (0, k) else + (apsnd \ times \ sgn) l (if sgn k = sgn l + then divmod_abs k l + else (let (r, s) = divmod_abs k l in + if s = 0 then (- r, 0) else (- r - 1, \l\ - s))))" +proof - + have aux1: "\k l::int. sgn k = sgn l \ k = 0 \ l = 0 \ 0 < l \ 0 < k \ l < 0 \ k < 0" + by (auto simp add: sgn_if) + have aux2: "\q::int. - int_of k = int_of l * q \ int_of k = int_of l * - q" by auto + show ?thesis + by (simp add: prod_eq_iff Target_Numeral.int_eq_iff prod_case_beta aux1) + (auto simp add: zdiv_zminus1_eq_if zmod_zminus1_eq_if zdiv_zminus2 zmod_zminus2 aux2) +qed + +lemma div_int_code [code]: + "k div l = fst (divmod k l)" + by simp + +lemma div_mod_code [code]: + "k mod l = snd (divmod k l)" + by simp + +lemma equal_int_code [code]: + "HOL.equal 0 (0::Target_Numeral.int) \ True" + "HOL.equal 0 (Pos l) \ False" + "HOL.equal 0 (Neg l) \ False" + "HOL.equal (Pos k) 0 \ False" + "HOL.equal (Pos k) (Pos l) \ HOL.equal k l" + "HOL.equal (Pos k) (Neg l) \ False" + "HOL.equal (Neg k) 0 \ False" + "HOL.equal (Neg k) (Pos l) \ False" + "HOL.equal (Neg k) (Neg l) \ HOL.equal k l" + by (simp_all add: equal Target_Numeral.int_eq_iff) + +lemma equal_int_refl [code nbe]: + "HOL.equal (k::Target_Numeral.int) k \ True" + by (fact equal_refl) + +lemma less_eq_int_code [code]: + "0 \ (0::Target_Numeral.int) \ True" + "0 \ Pos l \ True" + "0 \ Neg l \ False" + "Pos k \ 0 \ False" + "Pos k \ Pos l \ k \ l" + "Pos k \ Neg l \ False" + "Neg k \ 0 \ True" + "Neg k \ Pos l \ True" + "Neg k \ Neg l \ l \ k" + by (simp_all add: less_eq_int_def) + +lemma less_int_code [code]: + "0 < (0::Target_Numeral.int) \ False" + "0 < Pos l \ True" + "0 < Neg l \ False" + "Pos k < 0 \ False" + "Pos k < Pos l \ k < l" + "Pos k < Neg l \ False" + "Neg k < 0 \ True" + "Neg k < Pos l \ True" + "Neg k < Neg l \ l < k" + by (simp_all add: less_int_def) + +lemma nat_of_code [code]: + "nat_of (Neg k) = 0" + "nat_of 0 = 0" + "nat_of (Pos k) = nat_of_num k" + by (simp_all add: nat_of_def nat_of_num_numeral) + +lemma int_of_code [code]: + "int_of (Neg k) = neg_numeral k" + "int_of 0 = 0" + "int_of (Pos k) = numeral k" + by simp_all + +lemma of_int_code [code]: + "Target_Numeral.of_int (Int.Neg k) = neg_numeral k" + "Target_Numeral.of_int 0 = 0" + "Target_Numeral.of_int (Int.Pos k) = numeral k" + by simp_all + +definition num_of_int :: "Target_Numeral.int \ num" where + "num_of_int = num_of_nat \ nat_of" + +lemma num_of_int_code [code]: + "num_of_int k = (if k \ 1 then Num.One + else let + (l, j) = divmod k 2; + l' = num_of_int l + num_of_int l + in if j = 0 then l' else l' + Num.One)" +proof - + { + assume "int_of k mod 2 = 1" + then have "nat (int_of k mod 2) = nat 1" by simp + moreover assume *: "1 < int_of k" + ultimately have **: "nat (int_of k) mod 2 = 1" by (simp add: nat_mod_distrib) + have "num_of_nat (nat (int_of k)) = + num_of_nat (2 * (nat (int_of k) div 2) + nat (int_of k) mod 2)" + by simp + then have "num_of_nat (nat (int_of k)) = + num_of_nat (nat (int_of k) div 2 + nat (int_of k) div 2 + nat (int_of k) mod 2)" + by (simp add: nat_mult_2) + with ** have "num_of_nat (nat (int_of k)) = + num_of_nat (nat (int_of k) div 2 + nat (int_of k) div 2 + 1)" + by simp + } + note aux = this + show ?thesis + by (auto simp add: num_of_int_def nat_of_def Let_def prod_case_beta + not_le Target_Numeral.int_eq_iff less_eq_int_def + nat_mult_distrib nat_div_distrib num_of_nat_One num_of_nat_plus_distrib + nat_mult_2 aux add_One) +qed + +hide_const (open) int_of nat_of Pos Neg sub dup divmod_abs num_of_int + + +subsection {* Serializer setup for target language numerals *} + +code_type Target_Numeral.int + (SML "IntInf.int") + (OCaml "Big'_int.big'_int") + (Haskell "Integer") + (Scala "BigInt") + (Eval "int") + +code_instance Target_Numeral.int :: equal + (Haskell -) + +code_const "0::Target_Numeral.int" + (SML "0") + (OCaml "Big'_int.zero'_big'_int") + (Haskell "0") + (Scala "BigInt(0)") + +setup {* + fold (Numeral.add_code @{const_name Target_Numeral.Pos} + false Code_Printer.literal_numeral) ["SML", "OCaml", "Haskell", "Scala"] +*} + +setup {* + fold (Numeral.add_code @{const_name Target_Numeral.Neg} + true Code_Printer.literal_numeral) ["SML", "OCaml", "Haskell", "Scala"] +*} + +code_const "plus :: Target_Numeral.int \ _ \ _" + (SML "IntInf.+ ((_), (_))") + (OCaml "Big'_int.add'_big'_int") + (Haskell infixl 6 "+") + (Scala infixl 7 "+") + (Eval infixl 8 "+") + +code_const "uminus :: Target_Numeral.int \ _" + (SML "IntInf.~") + (OCaml "Big'_int.minus'_big'_int") + (Haskell "negate") + (Scala "!(- _)") + (Eval "~/ _") + +code_const "minus :: Target_Numeral.int \ _" + (SML "IntInf.- ((_), (_))") + (OCaml "Big'_int.sub'_big'_int") + (Haskell infixl 6 "-") + (Scala infixl 7 "-") + (Eval infixl 8 "-") + +code_const Target_Numeral.dup + (SML "IntInf.*/ (2,/ (_))") + (OCaml "Big'_int.mult'_big'_int/ 2") + (Haskell "!(2 * _)") + (Scala "!(2 * _)") + (Eval "!(2 * _)") + +code_const Target_Numeral.sub + (SML "!(raise/ Fail/ \"sub\")") + (OCaml "failwith/ \"sub\"") + (Haskell "error/ \"sub\"") + (Scala "!error(\"sub\")") + +code_const "times :: Target_Numeral.int \ _ \ _" + (SML "IntInf.* ((_), (_))") + (OCaml "Big'_int.mult'_big'_int") + (Haskell infixl 7 "*") + (Scala infixl 8 "*") + (Eval infixl 9 "*") + +code_const Target_Numeral.divmod_abs + (SML "IntInf.divMod/ (IntInf.abs _,/ IntInf.abs _)") + (OCaml "Big'_int.quomod'_big'_int/ (Big'_int.abs'_big'_int _)/ (Big'_int.abs'_big'_int _)") + (Haskell "divMod/ (abs _)/ (abs _)") + (Scala "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))") + (Eval "Integer.div'_mod/ (abs _)/ (abs _)") + +code_const "HOL.equal :: Target_Numeral.int \ _ \ bool" + (SML "!((_ : IntInf.int) = _)") + (OCaml "Big'_int.eq'_big'_int") + (Haskell infix 4 "==") + (Scala infixl 5 "==") + (Eval infixl 6 "=") + +code_const "less_eq :: Target_Numeral.int \ _ \ bool" + (SML "IntInf.<= ((_), (_))") + (OCaml "Big'_int.le'_big'_int") + (Haskell infix 4 "<=") + (Scala infixl 4 "<=") + (Eval infixl 6 "<=") + +code_const "less :: Target_Numeral.int \ _ \ bool" + (SML "IntInf.< ((_), (_))") + (OCaml "Big'_int.lt'_big'_int") + (Haskell infix 4 "<") + (Scala infixl 4 "<") + (Eval infixl 6 "<") + +ML {* +structure Target_Numeral = +struct + +val T = @{typ "Target_Numeral.int"}; + +end; +*} + +code_reserved Eval Target_Numeral + +code_const "Code_Evaluation.term_of \ Target_Numeral.int \ term" + (Eval "HOLogic.mk'_number/ Target'_Numeral.T") + +code_modulename SML + Target_Numeral Arith + +code_modulename OCaml + Target_Numeral Arith + +code_modulename Haskell + Target_Numeral Arith + + +subsection {* Implementation for @{typ int} *} + +code_datatype Target_Numeral.int_of + +lemma [code, code del]: + "Target_Numeral.of_int = Target_Numeral.of_int" .. + +lemma [code]: + "Target_Numeral.of_int (Target_Numeral.int_of k) = k" + by (simp add: Target_Numeral.int_eq_iff) + +declare Int.Pos_def [code] + +lemma [code_abbrev]: + "Target_Numeral.int_of (Target_Numeral.Pos k) = Int.Pos k" + by simp + +declare Int.Neg_def [code] + +lemma [code_abbrev]: + "Target_Numeral.int_of (Target_Numeral.Neg k) = Int.Neg k" + by simp + +lemma [code]: + "0 = Target_Numeral.int_of 0" + by simp + +lemma [code]: + "1 = Target_Numeral.int_of 1" + by simp + +lemma [code]: + "k + l = Target_Numeral.int_of (of_int k + of_int l)" + by simp + +lemma [code]: + "- k = Target_Numeral.int_of (- of_int k)" + by simp + +lemma [code]: + "k - l = Target_Numeral.int_of (of_int k - of_int l)" + by simp + +lemma [code]: + "Int.dup k = Target_Numeral.int_of (Target_Numeral.dup (of_int k))" + by simp + +lemma [code, code del]: + "Int.sub = Int.sub" .. + +lemma [code]: + "k * l = Target_Numeral.int_of (of_int k * of_int l)" + by simp + +lemma [code]: + "pdivmod k l = map_pair Target_Numeral.int_of Target_Numeral.int_of + (Target_Numeral.divmod_abs (of_int k) (of_int l))" + by (simp add: prod_eq_iff pdivmod_def) + +lemma [code]: + "k div l = Target_Numeral.int_of (of_int k div of_int l)" + by simp + +lemma [code]: + "k mod l = Target_Numeral.int_of (of_int k mod of_int l)" + by simp + +lemma [code]: + "HOL.equal k l = HOL.equal (of_int k :: Target_Numeral.int) (of_int l)" + by (simp add: equal Target_Numeral.int_eq_iff) + +lemma [code]: + "k \ l \ (of_int k :: Target_Numeral.int) \ of_int l" + by (simp add: less_eq_int_def) + +lemma [code]: + "k < l \ (of_int k :: Target_Numeral.int) < of_int l" + by (simp add: less_int_def) + +lemma (in ring_1) of_int_code: + "of_int k = (if k = 0 then 0 + else if k < 0 then - of_int (- k) + else let + (l, j) = divmod_int k 2; + l' = 2 * of_int l + in if j = 0 then l' else l' + 1)" +proof - + from mod_div_equality have *: "of_int k = of_int (k div 2 * 2 + k mod 2)" by simp + show ?thesis + by (simp add: Let_def divmod_int_mod_div mod_2_not_eq_zero_eq_one_int + of_int_add [symmetric]) (simp add: * mult_commute) +qed + +declare of_int_code [code] + + +subsection {* Implementation for @{typ nat} *} + +definition of_nat :: "nat \ Target_Numeral.int" where + [code_abbrev]: "of_nat = Nat.of_nat" + +hide_const (open) of_nat + +lemma int_of_nat [simp]: + "Target_Numeral.int_of (Target_Numeral.of_nat n) = of_nat n" + by (simp add: of_nat_def) + +lemma [code abstype]: + "Target_Numeral.nat_of (Target_Numeral.of_nat n) = n" + by (simp add: nat_of_def) + +lemma [code_abbrev]: + "nat (Int.Pos k) = nat_of_num k" + by (simp add: nat_of_num_numeral) + +lemma [code abstract]: + "Target_Numeral.of_nat 0 = 0" + by (simp add: Target_Numeral.int_eq_iff) + +lemma [code abstract]: + "Target_Numeral.of_nat 1 = 1" + by (simp add: Target_Numeral.int_eq_iff) + +lemma [code abstract]: + "Target_Numeral.of_nat (m + n) = of_nat m + of_nat n" + by (simp add: Target_Numeral.int_eq_iff) + +lemma [code abstract]: + "Target_Numeral.of_nat (Code_Nat.dup n) = Target_Numeral.dup (of_nat n)" + by (simp add: Target_Numeral.int_eq_iff Code_Nat.dup_def) + +lemma [code, code del]: + "Code_Nat.sub = Code_Nat.sub" .. + +lemma [code abstract]: + "Target_Numeral.of_nat (m - n) = max 0 (of_nat m - of_nat n)" + by (simp add: Target_Numeral.int_eq_iff) + +lemma [code abstract]: + "Target_Numeral.of_nat (m * n) = of_nat m * of_nat n" + by (simp add: Target_Numeral.int_eq_iff of_nat_mult) + +lemma [code abstract]: + "Target_Numeral.of_nat (m div n) = of_nat m div of_nat n" + by (simp add: Target_Numeral.int_eq_iff zdiv_int) + +lemma [code abstract]: + "Target_Numeral.of_nat (m mod n) = of_nat m mod of_nat n" + by (simp add: Target_Numeral.int_eq_iff zmod_int) + +lemma [code]: + "Divides.divmod_nat m n = (m div n, m mod n)" + by (simp add: prod_eq_iff) + +lemma [code]: + "HOL.equal m n = HOL.equal (of_nat m :: Target_Numeral.int) (of_nat n)" + by (simp add: equal Target_Numeral.int_eq_iff) + +lemma [code]: + "m \ n \ (of_nat m :: Target_Numeral.int) \ of_nat n" + by (simp add: less_eq_int_def) + +lemma [code]: + "m < n \ (of_nat m :: Target_Numeral.int) < of_nat n" + by (simp add: less_int_def) + +lemma num_of_nat_code [code]: + "num_of_nat = Target_Numeral.num_of_int \ Target_Numeral.of_nat" + by (simp add: fun_eq_iff num_of_int_def of_nat_def) + +lemma (in semiring_1) of_nat_code: + "of_nat n = (if n = 0 then 0 + else let + (m, q) = divmod_nat n 2; + m' = 2 * of_nat m + in if q = 0 then m' else m' + 1)" +proof - + from mod_div_equality have *: "of_nat n = of_nat (n div 2 * 2 + n mod 2)" by simp + show ?thesis + by (simp add: Let_def divmod_nat_div_mod mod_2_not_eq_zero_eq_one_nat + of_nat_add [symmetric]) + (simp add: * mult_commute of_nat_mult add_commute) +qed + +declare of_nat_code [code] + +text {* Conversions between @{typ nat} and @{typ int} *} + +definition int :: "nat \ int" where + [code_abbrev]: "int = of_nat" + +hide_const (open) int + +lemma [code]: + "Target_Numeral.int n = Target_Numeral.int_of (of_nat n)" + by (simp add: int_def) + +lemma [code abstract]: + "Target_Numeral.of_nat (nat k) = max 0 (Target_Numeral.of_int k)" + by (simp add: of_nat_def of_int_of_nat max_def) + +end diff -r 29e92b644d6c -r f067afe98049 src/HOL/List.thy --- a/src/HOL/List.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/List.thy Mon Mar 26 10:56:56 2012 +0200 @@ -2676,7 +2676,7 @@ -- {* simp does not terminate! *} by (induct j) auto -lemmas upt_rec_number_of[simp] = upt_rec[of "number_of m" "number_of n"] for m n +lemmas upt_rec_numeral[simp] = upt_rec[of "numeral m" "numeral n"] for m n lemma upt_conv_Nil [simp]: "j <= i ==> [i.. [i..j] = []" by(simp add: upto.simps) diff -r 29e92b644d6c -r f067afe98049 src/HOL/Matrix_LP/ComputeFloat.thy --- a/src/HOL/Matrix_LP/ComputeFloat.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Matrix_LP/ComputeFloat.thy Mon Mar 26 10:56:56 2012 +0200 @@ -75,8 +75,11 @@ ultimately show ?thesis by auto qed -lemma real_is_int_number_of[simp]: "real_is_int ((number_of \ int \ real) x)" - by (auto simp: real_is_int_def intro!: exI[of _ "number_of x"]) +lemma real_is_int_numeral[simp]: "real_is_int (numeral x)" + by (auto simp: real_is_int_def intro!: exI[of _ "numeral x"]) + +lemma real_is_int_neg_numeral[simp]: "real_is_int (neg_numeral x)" + by (auto simp: real_is_int_def intro!: exI[of _ "neg_numeral x"]) lemma int_of_real_0[simp]: "int_of_real (0::real) = (0::int)" by (simp add: int_of_real_def) @@ -87,7 +90,12 @@ show ?thesis by (simp only: 1 int_of_real_real) qed -lemma int_of_real_number_of[simp]: "int_of_real (number_of b) = number_of b" +lemma int_of_real_numeral[simp]: "int_of_real (numeral b) = numeral b" + unfolding int_of_real_def + by (intro some_equality) + (auto simp add: real_of_int_inject[symmetric] simp del: real_of_int_inject) + +lemma int_of_real_neg_numeral[simp]: "int_of_real (neg_numeral b) = neg_numeral b" unfolding int_of_real_def by (intro some_equality) (auto simp add: real_of_int_inject[symmetric] simp del: real_of_int_inject) @@ -101,7 +109,7 @@ lemma abs_div_2_less: "a \ 0 \ a \ -1 \ abs((a::int) div 2) < abs a" by arith -lemma norm_0_1: "(0::_::number_ring) = Numeral0 & (1::_::number_ring) = Numeral1" +lemma norm_0_1: "(1::_::numeral) = Numeral1" by auto lemma add_left_zero: "0 + a = (a::'a::comm_monoid_add)" @@ -116,34 +124,21 @@ lemma mult_right_one: "a * 1 = (a::'a::semiring_1)" by simp -lemma int_pow_0: "(a::int)^(Numeral0) = 1" +lemma int_pow_0: "(a::int)^0 = 1" by simp lemma int_pow_1: "(a::int)^(Numeral1) = a" by simp -lemma zero_eq_Numeral0_nring: "(0::'a::number_ring) = Numeral0" - by simp - -lemma one_eq_Numeral1_nring: "(1::'a::number_ring) = Numeral1" - by simp - -lemma zero_eq_Numeral0_nat: "(0::nat) = Numeral0" +lemma one_eq_Numeral1_nring: "(1::'a::numeral) = Numeral1" by simp lemma one_eq_Numeral1_nat: "(1::nat) = Numeral1" by simp -lemma zpower_Pls: "(z::int)^Numeral0 = Numeral1" +lemma zpower_Pls: "(z::int)^0 = Numeral1" by simp -lemma zpower_Min: "(z::int)^((-1)::nat) = Numeral1" -proof - - have 1:"((-1)::nat) = 0" - by simp - show ?thesis by (simp add: 1) -qed - lemma fst_cong: "a=a' \ fst (a,b) = fst (a',b)" by simp @@ -160,70 +155,8 @@ lemma not_true_eq_false: "(~ True) = False" by simp -lemmas binarith = - normalize_bin_simps - pred_bin_simps succ_bin_simps - add_bin_simps minus_bin_simps mult_bin_simps - -lemma int_eq_number_of_eq: - "(((number_of v)::int)=(number_of w)) = iszero ((number_of (v + uminus w))::int)" - by (rule eq_number_of_eq) - -lemma int_iszero_number_of_Pls: "iszero (Numeral0::int)" - by (simp only: iszero_number_of_Pls) - -lemma int_nonzero_number_of_Min: "~(iszero ((-1)::int))" - by simp - -lemma int_iszero_number_of_Bit0: "iszero ((number_of (Int.Bit0 w))::int) = iszero ((number_of w)::int)" - by simp - -lemma int_iszero_number_of_Bit1: "\ iszero ((number_of (Int.Bit1 w))::int)" - by simp - -lemma int_less_number_of_eq_neg: "(((number_of x)::int) < number_of y) = neg ((number_of (x + (uminus y)))::int)" - unfolding neg_def number_of_is_id by simp - -lemma int_not_neg_number_of_Pls: "\ (neg (Numeral0::int))" - by simp - -lemma int_neg_number_of_Min: "neg (-1::int)" - by simp - -lemma int_neg_number_of_Bit0: "neg ((number_of (Int.Bit0 w))::int) = neg ((number_of w)::int)" - by simp - -lemma int_neg_number_of_Bit1: "neg ((number_of (Int.Bit1 w))::int) = neg ((number_of w)::int)" - by simp - -lemma int_le_number_of_eq: "(((number_of x)::int) \ number_of y) = (\ neg ((number_of (y + (uminus x)))::int))" - unfolding neg_def number_of_is_id by (simp add: not_less) - -lemmas intarithrel = - int_eq_number_of_eq - lift_bool[OF int_iszero_number_of_Pls] nlift_bool[OF int_nonzero_number_of_Min] int_iszero_number_of_Bit0 - lift_bool[OF int_iszero_number_of_Bit1] int_less_number_of_eq_neg nlift_bool[OF int_not_neg_number_of_Pls] lift_bool[OF int_neg_number_of_Min] - int_neg_number_of_Bit0 int_neg_number_of_Bit1 int_le_number_of_eq - -lemma int_number_of_add_sym: "((number_of v)::int) + number_of w = number_of (v + w)" - by simp - -lemma int_number_of_diff_sym: "((number_of v)::int) - number_of w = number_of (v + (uminus w))" - by simp - -lemma int_number_of_mult_sym: "((number_of v)::int) * number_of w = number_of (v * w)" - by simp - -lemma int_number_of_minus_sym: "- ((number_of v)::int) = number_of (uminus v)" - by simp - -lemmas intarith = int_number_of_add_sym int_number_of_minus_sym int_number_of_diff_sym int_number_of_mult_sym - -lemmas natarith = add_nat_number_of diff_nat_number_of mult_nat_number_of eq_nat_number_of less_nat_number_of - -lemmas powerarith = nat_number_of zpower_number_of_even - zpower_number_of_odd[simplified zero_eq_Numeral0_nring one_eq_Numeral1_nring] - zpower_Pls zpower_Min +lemmas powerarith = nat_numeral zpower_numeral_even + zpower_numeral_odd zpower_Pls definition float :: "(int \ int) \ real" where "float = (\(a, b). real a * 2 powr real b)" @@ -302,7 +235,8 @@ float_minus float_abs zero_le_float float_pprt float_nprt pprt_lbound nprt_ubound (* for use with the compute oracle *) -lemmas arith = binarith intarith intarithrel natarith powerarith floatarith not_false_eq_true not_true_eq_false +lemmas arith = arith_simps rel_simps diff_nat_numeral nat_0 + nat_neg_numeral powerarith floatarith not_false_eq_true not_true_eq_false use "~~/src/HOL/Tools/float_arith.ML" diff -r 29e92b644d6c -r f067afe98049 src/HOL/Matrix_LP/ComputeNumeral.thy --- a/src/HOL/Matrix_LP/ComputeNumeral.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Matrix_LP/ComputeNumeral.thy Mon Mar 26 10:56:56 2012 +0200 @@ -2,145 +2,47 @@ imports ComputeHOL ComputeFloat begin -(* normalization of bit strings *) -lemmas bitnorm = normalize_bin_simps - -(* neg for bit strings *) -lemma neg1: "neg Int.Pls = False" by (simp add: Int.Pls_def) -lemma neg2: "neg Int.Min = True" apply (subst Int.Min_def) by auto -lemma neg3: "neg (Int.Bit0 x) = neg x" apply (simp add: neg_def) apply (subst Bit0_def) by auto -lemma neg4: "neg (Int.Bit1 x) = neg x" apply (simp add: neg_def) apply (subst Bit1_def) by auto -lemmas bitneg = neg1 neg2 neg3 neg4 - -(* iszero for bit strings *) -lemma iszero1: "iszero Int.Pls = True" by (simp add: Int.Pls_def iszero_def) -lemma iszero2: "iszero Int.Min = False" apply (subst Int.Min_def) apply (subst iszero_def) by simp -lemma iszero3: "iszero (Int.Bit0 x) = iszero x" apply (subst Int.Bit0_def) apply (subst iszero_def)+ by auto -lemma iszero4: "iszero (Int.Bit1 x) = False" apply (subst Int.Bit1_def) apply (subst iszero_def)+ apply simp by arith -lemmas bitiszero = iszero1 iszero2 iszero3 iszero4 - -(* lezero for bit strings *) -definition "lezero x \ x \ 0" -lemma lezero1: "lezero Int.Pls = True" unfolding Int.Pls_def lezero_def by auto -lemma lezero2: "lezero Int.Min = True" unfolding Int.Min_def lezero_def by auto -lemma lezero3: "lezero (Int.Bit0 x) = lezero x" unfolding Int.Bit0_def lezero_def by auto -lemma lezero4: "lezero (Int.Bit1 x) = neg x" unfolding Int.Bit1_def lezero_def neg_def by auto -lemmas bitlezero = lezero1 lezero2 lezero3 lezero4 - (* equality for bit strings *) -lemmas biteq = eq_bin_simps +lemmas biteq = eq_num_simps (* x < y for bit strings *) -lemmas bitless = less_bin_simps +lemmas bitless = less_num_simps (* x \ y for bit strings *) -lemmas bitle = le_bin_simps - -(* succ for bit strings *) -lemmas bitsucc = succ_bin_simps - -(* pred for bit strings *) -lemmas bitpred = pred_bin_simps - -(* unary minus for bit strings *) -lemmas bituminus = minus_bin_simps +lemmas bitle = le_num_simps (* addition for bit strings *) -lemmas bitadd = add_bin_simps +lemmas bitadd = add_num_simps (* multiplication for bit strings *) -lemma mult_Pls_right: "x * Int.Pls = Int.Pls" by (simp add: Pls_def) -lemma mult_Min_right: "x * Int.Min = - x" by (subst mult_commute) simp -lemma multb0x: "(Int.Bit0 x) * y = Int.Bit0 (x * y)" by (rule mult_Bit0) -lemma multxb0: "x * (Int.Bit0 y) = Int.Bit0 (x * y)" unfolding Bit0_def by simp -lemma multb1: "(Int.Bit1 x) * (Int.Bit1 y) = Int.Bit1 (Int.Bit0 (x * y) + x + y)" - unfolding Bit0_def Bit1_def by (simp add: algebra_simps) -lemmas bitmul = mult_Pls mult_Min mult_Pls_right mult_Min_right multb0x multxb0 multb1 +lemmas bitmul = mult_num_simps -lemmas bitarith = bitnorm bitiszero bitneg bitlezero biteq bitless bitle bitsucc bitpred bituminus bitadd bitmul - -definition "nat_norm_number_of (x::nat) = x" - -lemma nat_norm_number_of: "nat_norm_number_of (number_of w) = (if lezero w then 0 else number_of w)" - apply (simp add: nat_norm_number_of_def) - unfolding lezero_def iszero_def neg_def - apply (simp add: numeral_simps) - done +lemmas bitarith = arith_simps (* Normalization of nat literals *) -lemma natnorm0: "(0::nat) = number_of (Int.Pls)" by auto -lemma natnorm1: "(1 :: nat) = number_of (Int.Bit1 Int.Pls)" by auto -lemmas natnorm = natnorm0 natnorm1 nat_norm_number_of - -(* Suc *) -lemma natsuc: "Suc (number_of x) = (if neg x then 1 else number_of (Int.succ x))" by (auto simp add: number_of_is_id) - -(* Addition for nat *) -lemma natadd: "number_of x + ((number_of y)::nat) = (if neg x then (number_of y) else (if neg y then number_of x else (number_of (x + y))))" - unfolding nat_number_of_def number_of_is_id neg_def - by auto - -(* Subtraction for nat *) -lemma natsub: "(number_of x) - ((number_of y)::nat) = - (if neg x then 0 else (if neg y then number_of x else (nat_norm_number_of (number_of (x + (- y))))))" - unfolding nat_norm_number_of - by (auto simp add: number_of_is_id neg_def lezero_def iszero_def Let_def nat_number_of_def) - -(* Multiplication for nat *) -lemma natmul: "(number_of x) * ((number_of y)::nat) = - (if neg x then 0 else (if neg y then 0 else number_of (x * y)))" - unfolding nat_number_of_def number_of_is_id neg_def - by (simp add: nat_mult_distrib) - -lemma nateq: "(((number_of x)::nat) = (number_of y)) = ((lezero x \ lezero y) \ (x = y))" - by (auto simp add: iszero_def lezero_def neg_def number_of_is_id) - -lemma natless: "(((number_of x)::nat) < (number_of y)) = ((x < y) \ (\ (lezero y)))" - by (simp add: lezero_def numeral_simps not_le) - -lemma natle: "(((number_of x)::nat) \ (number_of y)) = (y < x \ lezero x)" - by (auto simp add: number_of_is_id lezero_def nat_number_of_def) +lemmas natnorm = one_eq_Numeral1_nat fun natfac :: "nat \ nat" where "natfac n = (if n = 0 then 1 else n * (natfac (n - 1)))" -lemmas compute_natarith = bitarith natnorm natsuc natadd natsub natmul nateq natless natle natfac.simps - -lemma number_eq: "(((number_of x)::'a::{number_ring, linordered_idom}) = (number_of y)) = (x = y)" - unfolding number_of_eq - apply simp - done +lemmas compute_natarith = + arith_simps rel_simps + diff_nat_numeral nat_numeral nat_0 nat_neg_numeral + numeral_1_eq_1 [symmetric] + numeral_1_eq_Suc_0 [symmetric] + Suc_numeral natfac.simps -lemma number_le: "(((number_of x)::'a::{number_ring, linordered_idom}) \ (number_of y)) = (x \ y)" - unfolding number_of_eq - apply simp - done - -lemma number_less: "(((number_of x)::'a::{number_ring, linordered_idom}) < (number_of y)) = (x < y)" - unfolding number_of_eq - apply simp - done +lemmas number_norm = numeral_1_eq_1[symmetric] -lemma number_diff: "((number_of x)::'a::{number_ring, linordered_idom}) - number_of y = number_of (x + (- y))" - apply (subst diff_number_of_eq) - apply simp - done - -lemmas number_norm = number_of_Pls[symmetric] numeral_1_eq_1[symmetric] - -lemmas compute_numberarith = number_of_minus[symmetric] number_of_add[symmetric] number_diff number_of_mult[symmetric] number_norm number_eq number_le number_less +lemmas compute_numberarith = + arith_simps rel_simps number_norm -lemma compute_real_of_nat_number_of: "real ((number_of v)::nat) = (if neg v then 0 else number_of v)" - by (simp only: real_of_nat_number_of number_of_is_id) - -lemma compute_nat_of_int_number_of: "nat ((number_of v)::int) = (number_of v)" - by simp +lemmas compute_num_conversions = + real_of_nat_numeral real_of_nat_zero + nat_numeral nat_0 nat_neg_numeral + real_numeral real_of_int_zero -lemmas compute_num_conversions = compute_real_of_nat_number_of compute_nat_of_int_number_of real_number_of - -lemmas zpowerarith = zpower_number_of_even - zpower_number_of_odd[simplified zero_eq_Numeral0_nring one_eq_Numeral1_nring] - zpower_Pls zpower_Min +lemmas zpowerarith = zpower_numeral_even zpower_numeral_odd zpower_Pls int_pow_1 (* div, mod *) @@ -162,26 +64,19 @@ (* collecting all the theorems *) -lemma even_Pls: "even (Int.Pls) = True" - apply (unfold Pls_def even_def) +lemma even_0_int: "even (0::int) = True" by simp -lemma even_Min: "even (Int.Min) = False" - apply (unfold Min_def even_def) +lemma even_One_int: "even (numeral Num.One :: int) = False" by simp -lemma even_B0: "even (Int.Bit0 x) = True" - apply (unfold Bit0_def) +lemma even_Bit0_int: "even (numeral (Num.Bit0 x) :: int) = True" by simp -lemma even_B1: "even (Int.Bit1 x) = False" - apply (unfold Bit1_def) +lemma even_Bit1_int: "even (numeral (Num.Bit1 x) :: int) = False" by simp -lemma even_number_of: "even ((number_of w)::int) = even w" - by (simp only: number_of_is_id) - -lemmas compute_even = even_Pls even_Min even_B0 even_B1 even_number_of +lemmas compute_even = even_0_int even_One_int even_Bit0_int even_Bit1_int lemmas compute_numeral = compute_if compute_let compute_pair compute_bool compute_natarith compute_numberarith max_def min_def compute_num_conversions zpowerarith compute_div_mod compute_even diff -r 29e92b644d6c -r f067afe98049 src/HOL/Matrix_LP/SparseMatrix.thy --- a/src/HOL/Matrix_LP/SparseMatrix.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Matrix_LP/SparseMatrix.thy Mon Mar 26 10:56:56 2012 +0200 @@ -1029,9 +1029,7 @@ sparse_row_matrix_pprt sorted_spvec_pprt_spmat sorted_spmat_pprt_spmat sparse_row_matrix_nprt sorted_spvec_nprt_spmat sorted_spmat_nprt_spmat -lemma zero_eq_Numeral0: "(0::_::number_ring) = Numeral0" by simp - -lemmas sparse_row_matrix_arith_simps[simplified zero_eq_Numeral0] = +lemmas sparse_row_matrix_arith_simps = mult_spmat.simps mult_spvec_spmat.simps addmult_spvec.simps smult_spvec_empty smult_spvec_cons diff -r 29e92b644d6c -r f067afe98049 src/HOL/Metis_Examples/Big_O.thy --- a/src/HOL/Metis_Examples/Big_O.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Metis_Examples/Big_O.thy Mon Mar 26 10:56:56 2012 +0200 @@ -16,7 +16,7 @@ subsection {* Definitions *} -definition bigo :: "('a => 'b\{linordered_idom,number_ring}) => ('a => 'b) set" ("(1O'(_'))") where +definition bigo :: "('a => 'b\linordered_idom) => ('a => 'b) set" ("(1O'(_'))") where "O(f\('a => 'b)) == {h. \c. \x. abs (h x) <= c * abs (f x)}" lemma bigo_pos_const: @@ -180,7 +180,7 @@ apply (rule_tac x = "c + c" in exI) apply auto apply (subgoal_tac "c * abs (f xa + g xa) <= (c + c) * abs (g xa)") - apply (metis order_trans semiring_mult_2) + apply (metis order_trans mult_2) apply (subgoal_tac "c * abs (f xa + g xa) <= c * (abs (f xa) + abs (g xa))") apply (erule order_trans) apply (simp add: ring_distribs) @@ -325,7 +325,7 @@ by (metis bigo_mult2 set_plus_mono_b set_times_intro2 set_times_plus_distrib) lemma bigo_mult5: "\x. f x ~= 0 \ - O(f * g) <= (f\'a => ('b\{linordered_field,number_ring})) *o O(g)" + O(f * g) <= (f\'a => ('b\linordered_field)) *o O(g)" proof - assume a: "\x. f x ~= 0" show "O(f * g) <= f *o O(g)" @@ -351,21 +351,21 @@ qed lemma bigo_mult6: -"\x. f x \ 0 \ O(f * g) = (f\'a \ ('b\{linordered_field,number_ring})) *o O(g)" +"\x. f x \ 0 \ O(f * g) = (f\'a \ ('b\linordered_field)) *o O(g)" by (metis bigo_mult2 bigo_mult5 order_antisym) (*proof requires relaxing relevance: 2007-01-25*) declare bigo_mult6 [simp] lemma bigo_mult7: -"\x. f x \ 0 \ O(f * g) \ O(f\'a \ ('b\{linordered_field,number_ring})) \ O(g)" +"\x. f x \ 0 \ O(f * g) \ O(f\'a \ ('b\linordered_field)) \ O(g)" by (metis bigo_refl bigo_mult6 set_times_mono3) declare bigo_mult6 [simp del] declare bigo_mult7 [intro!] lemma bigo_mult8: -"\x. f x \ 0 \ O(f * g) = O(f\'a \ ('b\{linordered_field,number_ring})) \ O(g)" +"\x. f x \ 0 \ O(f * g) = O(f\'a \ ('b\linordered_field)) \ O(g)" by (metis bigo_mult bigo_mult7 order_antisym_conv) lemma bigo_minus [intro]: "f : O(g) \ - f : O(g)" @@ -405,14 +405,14 @@ lemma bigo_const2 [intro]: "O(\x. c) \ O(\x. 1)" by (metis bigo_const1 bigo_elt_subset) -lemma bigo_const3: "(c\'a\{linordered_field,number_ring}) ~= 0 \ (\x. 1) : O(\x. c)" +lemma bigo_const3: "(c\'a\linordered_field) ~= 0 \ (\x. 1) : O(\x. c)" apply (simp add: bigo_def) by (metis abs_eq_0 left_inverse order_refl) -lemma bigo_const4: "(c\'a\{linordered_field,number_ring}) ~= 0 \ O(\x. 1) <= O(\x. c)" +lemma bigo_const4: "(c\'a\linordered_field) ~= 0 \ O(\x. 1) <= O(\x. c)" by (metis bigo_elt_subset bigo_const3) -lemma bigo_const [simp]: "(c\'a\{linordered_field,number_ring}) ~= 0 \ +lemma bigo_const [simp]: "(c\'a\linordered_field) ~= 0 \ O(\x. c) = O(\x. 1)" by (metis bigo_const2 bigo_const4 equalityI) @@ -423,19 +423,19 @@ lemma bigo_const_mult2: "O(\x. c * f x) \ O(f)" by (rule bigo_elt_subset, rule bigo_const_mult1) -lemma bigo_const_mult3: "(c\'a\{linordered_field,number_ring}) ~= 0 \ f : O(\x. c * f x)" +lemma bigo_const_mult3: "(c\'a\linordered_field) ~= 0 \ f : O(\x. c * f x)" apply (simp add: bigo_def) by (metis (no_types) abs_mult mult_assoc mult_1 order_refl left_inverse) lemma bigo_const_mult4: -"(c\'a\{linordered_field,number_ring}) \ 0 \ O(f) \ O(\x. c * f x)" +"(c\'a\linordered_field) \ 0 \ O(f) \ O(\x. c * f x)" by (metis bigo_elt_subset bigo_const_mult3) -lemma bigo_const_mult [simp]: "(c\'a\{linordered_field,number_ring}) ~= 0 \ +lemma bigo_const_mult [simp]: "(c\'a\linordered_field) ~= 0 \ O(\x. c * f x) = O(f)" by (metis equalityI bigo_const_mult2 bigo_const_mult4) -lemma bigo_const_mult5 [simp]: "(c\'a\{linordered_field,number_ring}) ~= 0 \ +lemma bigo_const_mult5 [simp]: "(c\'a\linordered_field) ~= 0 \ (\x. c) *o O(f) = O(f)" apply (auto del: subsetI) apply (rule order_trans) @@ -587,7 +587,7 @@ apply assumption+ done -lemma bigo_useful_const_mult: "(c\'a\{linordered_field,number_ring}) ~= 0 \ +lemma bigo_useful_const_mult: "(c\'a\linordered_field) ~= 0 \ (\x. c) * f =o O(h) \ f =o O(h)" apply (rule subsetD) apply (subgoal_tac "(\x. 1 / c) *o O(h) <= O(h)") @@ -696,7 +696,7 @@ by (metis abs_ge_zero linorder_linear min_max.sup_absorb1 min_max.sup_commute) lemma bigo_lesso4: - "f 'a=>'b\{linordered_field,number_ring}) \ + "f 'a=>'b\{linordered_field}) \ g =o h +o O(k) \ f \(x::'a::euclidean_space). norm x = c" apply (rule exI[where x="c *\<^sub>R basis 0 ::'a"]) using DIM_positive[where 'a='a] by auto diff -r 29e92b644d6c -r f067afe98049 src/HOL/Multivariate_Analysis/Determinants.thy --- a/src/HOL/Multivariate_Analysis/Determinants.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Multivariate_Analysis/Determinants.thy Mon Mar 26 10:56:56 2012 +0200 @@ -286,7 +286,7 @@ proof- have tha: "\(a::'a) b. a = b ==> b = - a ==> a = 0" by simp - have th1: "of_int (-1) = - 1" by (metis of_int_1 of_int_minus number_of_Min) + have th1: "of_int (-1) = - 1" by simp let ?p = "Fun.swap i j id" let ?A = "\ i. A $ ?p i" from r have "A = ?A" by (simp add: vec_eq_iff row_def swap_def) @@ -1058,8 +1058,7 @@ unfolding det_def UNIV_2 unfolding setsum_over_permutations_insert[OF f12] unfolding permutes_sing - apply (simp add: sign_swap_id sign_id swap_id_eq) - by (simp add: arith_simps(31)[symmetric] del: arith_simps(31)) + by (simp add: sign_swap_id sign_id swap_id_eq) qed lemma det_3: "det (A::'a::comm_ring_1^3^3) = @@ -1079,9 +1078,7 @@ unfolding setsum_over_permutations_insert[OF f23] unfolding permutes_sing - apply (simp add: sign_swap_id permutation_swap_id sign_compose sign_id swap_id_eq) - apply (simp add: arith_simps(31)[symmetric] del: arith_simps(31)) - by (simp add: field_simps) + by (simp add: sign_swap_id permutation_swap_id sign_compose sign_id swap_id_eq) qed end diff -r 29e92b644d6c -r f067afe98049 src/HOL/Multivariate_Analysis/Norm_Arith.thy --- a/src/HOL/Multivariate_Analysis/Norm_Arith.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Multivariate_Analysis/Norm_Arith.thy Mon Mar 26 10:56:56 2012 +0200 @@ -104,6 +104,17 @@ "x \ y \ \ (norm (x - y) \ 0)" using norm_ge_zero[of "x - y"] by auto +lemmas arithmetic_simps = + arith_simps + add_numeral_special + add_neg_numeral_special + add_0_left + add_0_right + mult_zero_left + mult_zero_right + mult_1_left + mult_1_right + use "normarith.ML" method_setup norm = {* Scan.succeed (SIMPLE_METHOD' o NormArith.norm_arith_tac) diff -r 29e92b644d6c -r f067afe98049 src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Mar 26 10:56:56 2012 +0200 @@ -5786,7 +5786,7 @@ { assume as:"dist a b > dist (f n x) (f n y)" then obtain Na Nb where "\m\Na. dist (f (r m) x) a < (dist a b - dist (f n x) (f n y)) / 2" and "\m\Nb. dist (f (r m) y) b < (dist a b - dist (f n x) (f n y)) / 2" - using lima limb unfolding h_def LIMSEQ_def by (fastforce simp del: less_divide_eq_number_of1) + using lima limb unfolding h_def LIMSEQ_def by (fastforce simp del: less_divide_eq_numeral1) hence "dist (f (r (Na + Nb + n)) x - f (r (Na + Nb + n)) y) (a - b) < dist a b - dist (f n x) (f n y)" apply(erule_tac x="Na+Nb+n" in allE) apply(erule_tac x="Na+Nb+n" in allE) apply simp diff -r 29e92b644d6c -r f067afe98049 src/HOL/Mutabelle/mutabelle_extra.ML --- a/src/HOL/Mutabelle/mutabelle_extra.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Mutabelle/mutabelle_extra.ML Mon Mar 26 10:56:56 2012 +0200 @@ -271,7 +271,7 @@ @{const_name enum_prod_inst.enum_ex_prod}, @{const_name Quickcheck.catch_match}, @{const_name Quickcheck_Exhaustive.unknown}, - @{const_name Int.Bit0}, @{const_name Int.Bit1} + @{const_name Num.Bit0}, @{const_name Num.Bit1} (*@{const_name "==>"}, @{const_name "=="}*)] val forbidden_mutant_consts = diff -r 29e92b644d6c -r f067afe98049 src/HOL/NSA/HyperDef.thy --- a/src/HOL/NSA/HyperDef.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/NSA/HyperDef.thy Mon Mar 26 10:56:56 2012 +0200 @@ -346,8 +346,8 @@ K (Lin_Arith.add_inj_thms [@{thm star_of_le} RS iffD2, @{thm star_of_less} RS iffD2, @{thm star_of_eq} RS iffD2] #> Lin_Arith.add_simps [@{thm star_of_zero}, @{thm star_of_one}, - @{thm star_of_number_of}, @{thm star_of_add}, @{thm star_of_minus}, - @{thm star_of_diff}, @{thm star_of_mult}] + @{thm star_of_numeral}, @{thm star_of_neg_numeral}, @{thm star_of_add}, + @{thm star_of_minus}, @{thm star_of_diff}, @{thm star_of_mult}] #> Lin_Arith.add_inj_const (@{const_name "StarDef.star_of"}, @{typ "real \ hypreal"})) *} @@ -419,10 +419,15 @@ x ^ Suc (Suc 0) + y ^ Suc (Suc 0) + (hypreal_of_nat (Suc (Suc 0)))*x*y" by (simp add: right_distrib left_distrib) -lemma power_hypreal_of_real_number_of: - "(number_of v :: hypreal) ^ n = hypreal_of_real ((number_of v) ^ n)" +lemma power_hypreal_of_real_numeral: + "(numeral v :: hypreal) ^ n = hypreal_of_real ((numeral v) ^ n)" by simp -declare power_hypreal_of_real_number_of [of _ "number_of w", simp] for w +declare power_hypreal_of_real_numeral [of _ "numeral w", simp] for w + +lemma power_hypreal_of_real_neg_numeral: + "(neg_numeral v :: hypreal) ^ n = hypreal_of_real ((neg_numeral v) ^ n)" +by simp +declare power_hypreal_of_real_neg_numeral [of _ "numeral w", simp] for w (* lemma hrealpow_HFinite: fixes x :: "'a::{real_normed_algebra,power} star" @@ -492,7 +497,7 @@ by transfer (rule power_one) lemma hrabs_hyperpow_minus_one [simp]: - "\n. abs(-1 pow n) = (1::'a::{number_ring,linordered_idom} star)" + "\n. abs(-1 pow n) = (1::'a::{linordered_idom} star)" by transfer (rule abs_power_minus_one) lemma hyperpow_mult: diff -r 29e92b644d6c -r f067afe98049 src/HOL/NSA/NSA.thy --- a/src/HOL/NSA/NSA.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/NSA/NSA.thy Mon Mar 26 10:56:56 2012 +0200 @@ -190,7 +190,7 @@ lemma SReal_hypreal_of_real [simp]: "hypreal_of_real x \ Reals" by (simp add: Reals_eq_Standard) -lemma SReal_divide_number_of: "r \ Reals ==> r/(number_of w::hypreal) \ Reals" +lemma SReal_divide_numeral: "r \ Reals ==> r/(numeral w::hypreal) \ Reals" by simp text{*epsilon is not in Reals because it is an infinitesimal*} @@ -290,8 +290,8 @@ "(hnorm (x::hypreal) \ HFinite) = (x \ HFinite)" by (simp add: HFinite_def) -lemma HFinite_number_of [simp]: "number_of w \ HFinite" -unfolding star_number_def by (rule HFinite_star_of) +lemma HFinite_numeral [simp]: "numeral w \ HFinite" +unfolding star_numeral_def by (rule HFinite_star_of) (** As always with numerals, 0 and 1 are special cases **) @@ -347,7 +347,7 @@ apply (rule InfinitesimalI) apply (rule hypreal_sum_of_halves [THEN subst]) apply (drule half_gt_zero) -apply (blast intro: hnorm_add_less SReal_divide_number_of dest: InfinitesimalD) +apply (blast intro: hnorm_add_less SReal_divide_numeral dest: InfinitesimalD) done lemma Infinitesimal_minus_iff [simp]: "(-x:Infinitesimal) = (x:Infinitesimal)" @@ -652,7 +652,7 @@ (*reorientation simplification procedure: reorients (polymorphic) 0 = x, 1 = x, nnn = x provided x isn't 0, 1 or a numeral.*) simproc_setup approx_reorient_simproc - ("0 @= x" | "1 @= y" | "number_of w @= z") = + ("0 @= x" | "1 @= y" | "numeral w @= z" | "neg_numeral w @= r") = {* let val rule = @{thm approx_reorient} RS eq_reflection fun proc phi ss ct = case term_of ct of @@ -957,9 +957,9 @@ "x \ 0 ==> star_of x \ HFinite - Infinitesimal" by simp -lemma number_of_not_Infinitesimal [simp]: - "number_of w \ (0::hypreal) ==> (number_of w :: hypreal) \ Infinitesimal" -by (fast dest: Reals_number_of [THEN SReal_Infinitesimal_zero]) +lemma numeral_not_Infinitesimal [simp]: + "numeral w \ (0::hypreal) ==> (numeral w :: hypreal) \ Infinitesimal" +by (fast dest: Reals_numeral [THEN SReal_Infinitesimal_zero]) (*again: 1 is a special case, but not 0 this time*) lemma one_not_Infinitesimal [simp]: @@ -1024,31 +1024,31 @@ apply simp done -lemma number_of_approx_iff [simp]: - "(number_of v @= (number_of w :: 'a::{number,real_normed_vector} star)) = - (number_of v = (number_of w :: 'a))" -apply (unfold star_number_def) +lemma numeral_approx_iff [simp]: + "(numeral v @= (numeral w :: 'a::{numeral,real_normed_vector} star)) = + (numeral v = (numeral w :: 'a))" +apply (unfold star_numeral_def) apply (rule star_of_approx_iff) done (*And also for 0 @= #nn and 1 @= #nn, #nn @= 0 and #nn @= 1.*) lemma [simp]: - "(number_of w @= (0::'a::{number,real_normed_vector} star)) = - (number_of w = (0::'a))" - "((0::'a::{number,real_normed_vector} star) @= number_of w) = - (number_of w = (0::'a))" - "(number_of w @= (1::'b::{number,one,real_normed_vector} star)) = - (number_of w = (1::'b))" - "((1::'b::{number,one,real_normed_vector} star) @= number_of w) = - (number_of w = (1::'b))" + "(numeral w @= (0::'a::{numeral,real_normed_vector} star)) = + (numeral w = (0::'a))" + "((0::'a::{numeral,real_normed_vector} star) @= numeral w) = + (numeral w = (0::'a))" + "(numeral w @= (1::'b::{numeral,one,real_normed_vector} star)) = + (numeral w = (1::'b))" + "((1::'b::{numeral,one,real_normed_vector} star) @= numeral w) = + (numeral w = (1::'b))" "~ (0 @= (1::'c::{zero_neq_one,real_normed_vector} star))" "~ (1 @= (0::'c::{zero_neq_one,real_normed_vector} star))" -apply (unfold star_number_def star_zero_def star_one_def) +apply (unfold star_numeral_def star_zero_def star_one_def) apply (unfold star_of_approx_iff) by (auto intro: sym) -lemma star_of_approx_number_of_iff [simp]: - "(star_of k @= number_of w) = (k = number_of w)" +lemma star_of_approx_numeral_iff [simp]: + "(star_of k @= numeral w) = (k = numeral w)" by (subst star_of_approx_iff [symmetric], auto) lemma star_of_approx_zero_iff [simp]: "(star_of k @= 0) = (k = 0)" @@ -1843,8 +1843,11 @@ lemma st_add: "\x \ HFinite; y \ HFinite\ \ st (x + y) = st x + st y" by (simp add: st_unique st_SReal st_approx_self approx_add) -lemma st_number_of [simp]: "st (number_of w) = number_of w" -by (rule Reals_number_of [THEN st_SReal_eq]) +lemma st_numeral [simp]: "st (numeral w) = numeral w" +by (rule Reals_numeral [THEN st_SReal_eq]) + +lemma st_neg_numeral [simp]: "st (neg_numeral w) = neg_numeral w" +by (rule Reals_neg_numeral [THEN st_SReal_eq]) lemma st_0 [simp]: "st 0 = 0" by (simp add: st_SReal_eq) diff -r 29e92b644d6c -r f067afe98049 src/HOL/NSA/NSCA.thy --- a/src/HOL/NSA/NSCA.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/NSA/NSCA.thy Mon Mar 26 10:56:56 2012 +0200 @@ -32,14 +32,14 @@ "hcmod (hcomplex_of_complex r) \ Reals" by (simp add: Reals_eq_Standard) -lemma SReal_hcmod_number_of [simp]: "hcmod (number_of w ::hcomplex) \ Reals" +lemma SReal_hcmod_numeral [simp]: "hcmod (numeral w ::hcomplex) \ Reals" by (simp add: Reals_eq_Standard) lemma SReal_hcmod_SComplex: "x \ SComplex ==> hcmod x \ Reals" by (simp add: Reals_eq_Standard) -lemma SComplex_divide_number_of: - "r \ SComplex ==> r/(number_of w::hcomplex) \ SComplex" +lemma SComplex_divide_numeral: + "r \ SComplex ==> r/(numeral w::hcomplex) \ SComplex" by simp lemma SComplex_UNIV_complex: @@ -211,9 +211,9 @@ ==> hcomplex_of_complex x \ HFinite - Infinitesimal" by (rule SComplex_HFinite_diff_Infinitesimal, auto) -lemma number_of_not_Infinitesimal [simp]: - "number_of w \ (0::hcomplex) ==> (number_of w::hcomplex) \ Infinitesimal" -by (fast dest: Standard_number_of [THEN SComplex_Infinitesimal_zero]) +lemma numeral_not_Infinitesimal [simp]: + "numeral w \ (0::hcomplex) ==> (numeral w::hcomplex) \ Infinitesimal" +by (fast dest: Standard_numeral [THEN SComplex_Infinitesimal_zero]) lemma approx_SComplex_not_zero: "[| y \ SComplex; x @= y; y\ 0 |] ==> x \ 0" @@ -223,11 +223,11 @@ "[|x \ SComplex; y \ SComplex|] ==> (x @= y) = (x = y)" by (auto simp add: Standard_def) -lemma number_of_Infinitesimal_iff [simp]: - "((number_of w :: hcomplex) \ Infinitesimal) = - (number_of w = (0::hcomplex))" +lemma numeral_Infinitesimal_iff [simp]: + "((numeral w :: hcomplex) \ Infinitesimal) = + (numeral w = (0::hcomplex))" apply (rule iffI) -apply (fast dest: Standard_number_of [THEN SComplex_Infinitesimal_zero]) +apply (fast dest: Standard_numeral [THEN SComplex_Infinitesimal_zero]) apply (simp (no_asm_simp)) done @@ -441,8 +441,8 @@ "[| x \ HFinite; y \ HFinite |] ==> stc (x + y) = stc(x) + stc(y)" by (simp add: stc_unique stc_SComplex stc_approx_self approx_add) -lemma stc_number_of [simp]: "stc (number_of w) = number_of w" -by (rule Standard_number_of [THEN stc_SComplex_eq]) +lemma stc_numeral [simp]: "stc (numeral w) = numeral w" +by (rule Standard_numeral [THEN stc_SComplex_eq]) lemma stc_zero [simp]: "stc 0 = 0" by simp diff -r 29e92b644d6c -r f067afe98049 src/HOL/NSA/NSComplex.thy --- a/src/HOL/NSA/NSComplex.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/NSA/NSComplex.thy Mon Mar 26 10:56:56 2012 +0200 @@ -626,32 +626,38 @@ subsection{*Numerals and Arithmetic*} -lemma hcomplex_number_of_def: "(number_of w :: hcomplex) == of_int w" -by transfer (rule number_of_eq [THEN eq_reflection]) - lemma hcomplex_of_hypreal_eq_hcomplex_of_complex: "hcomplex_of_hypreal (hypreal_of_real x) = hcomplex_of_complex (complex_of_real x)" by transfer (rule refl) -lemma hcomplex_hypreal_number_of: - "hcomplex_of_complex (number_of w) = hcomplex_of_hypreal(number_of w)" -by transfer (rule of_real_number_of_eq [symmetric]) +lemma hcomplex_hypreal_numeral: + "hcomplex_of_complex (numeral w) = hcomplex_of_hypreal(numeral w)" +by transfer (rule of_real_numeral [symmetric]) -lemma hcomplex_number_of_hcnj [simp]: - "hcnj (number_of v :: hcomplex) = number_of v" -by transfer (rule complex_cnj_number_of) +lemma hcomplex_hypreal_neg_numeral: + "hcomplex_of_complex (neg_numeral w) = hcomplex_of_hypreal(neg_numeral w)" +by transfer (rule of_real_neg_numeral [symmetric]) + +lemma hcomplex_numeral_hcnj [simp]: + "hcnj (numeral v :: hcomplex) = numeral v" +by transfer (rule complex_cnj_numeral) -lemma hcomplex_number_of_hcmod [simp]: - "hcmod(number_of v :: hcomplex) = abs (number_of v :: hypreal)" -by transfer (rule norm_number_of) +lemma hcomplex_numeral_hcmod [simp]: + "hcmod(numeral v :: hcomplex) = (numeral v :: hypreal)" +by transfer (rule norm_numeral) + +lemma hcomplex_neg_numeral_hcmod [simp]: + "hcmod(neg_numeral v :: hcomplex) = (numeral v :: hypreal)" +by transfer (rule norm_neg_numeral) -lemma hcomplex_number_of_hRe [simp]: - "hRe(number_of v :: hcomplex) = number_of v" -by transfer (rule complex_Re_number_of) +lemma hcomplex_numeral_hRe [simp]: + "hRe(numeral v :: hcomplex) = numeral v" +by transfer (rule complex_Re_numeral) -lemma hcomplex_number_of_hIm [simp]: - "hIm(number_of v :: hcomplex) = 0" -by transfer (rule complex_Im_number_of) +lemma hcomplex_numeral_hIm [simp]: + "hIm(numeral v :: hcomplex) = 0" +by transfer (rule complex_Im_numeral) +(* TODO: add neg_numeral rules above *) end diff -r 29e92b644d6c -r f067afe98049 src/HOL/NSA/StarDef.thy --- a/src/HOL/NSA/StarDef.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/NSA/StarDef.thy Mon Mar 26 10:56:56 2012 +0200 @@ -522,16 +522,6 @@ end -instantiation star :: (number) number -begin - -definition - star_number_def: "number_of b \ star_of (number_of b)" - -instance .. - -end - instance star :: (Rings.dvd) Rings.dvd .. instantiation star :: (Divides.div) Divides.div @@ -561,7 +551,7 @@ end lemmas star_class_defs [transfer_unfold] = - star_zero_def star_one_def star_number_def + star_zero_def star_one_def star_add_def star_diff_def star_minus_def star_mult_def star_divide_def star_inverse_def star_le_def star_less_def star_abs_def star_sgn_def @@ -575,9 +565,6 @@ lemma Standard_one: "1 \ Standard" by (simp add: star_one_def) -lemma Standard_number_of: "number_of b \ Standard" -by (simp add: star_number_def) - lemma Standard_add: "\x \ Standard; y \ Standard\ \ x + y \ Standard" by (simp add: star_add_def) @@ -606,7 +593,7 @@ by (simp add: star_mod_def) lemmas Standard_simps [simp] = - Standard_zero Standard_one Standard_number_of + Standard_zero Standard_one Standard_add Standard_diff Standard_minus Standard_mult Standard_divide Standard_inverse Standard_abs Standard_div Standard_mod @@ -648,9 +635,6 @@ lemma star_of_one: "star_of 1 = 1" by transfer (rule refl) -lemma star_of_number_of: "star_of (number_of x) = number_of x" -by transfer (rule refl) - text {* @{term star_of} preserves orderings *} lemma star_of_less: "(star_of x < star_of y) = (x < y)" @@ -682,34 +666,16 @@ lemmas star_of_le_1 = star_of_le [of _ 1, simplified star_of_one] lemmas star_of_eq_1 = star_of_eq [of _ 1, simplified star_of_one] -text{*As above, for numerals*} - -lemmas star_of_number_less = - star_of_less [of "number_of w", simplified star_of_number_of] for w -lemmas star_of_number_le = - star_of_le [of "number_of w", simplified star_of_number_of] for w -lemmas star_of_number_eq = - star_of_eq [of "number_of w", simplified star_of_number_of] for w - -lemmas star_of_less_number = - star_of_less [of _ "number_of w", simplified star_of_number_of] for w -lemmas star_of_le_number = - star_of_le [of _ "number_of w", simplified star_of_number_of] for w -lemmas star_of_eq_number = - star_of_eq [of _ "number_of w", simplified star_of_number_of] for w - lemmas star_of_simps [simp] = star_of_add star_of_diff star_of_minus star_of_mult star_of_divide star_of_inverse star_of_div star_of_mod star_of_abs - star_of_zero star_of_one star_of_number_of + star_of_zero star_of_one star_of_less star_of_le star_of_eq star_of_0_less star_of_0_le star_of_0_eq star_of_less_0 star_of_le_0 star_of_eq_0 star_of_1_less star_of_1_le star_of_1_eq star_of_less_1 star_of_le_1 star_of_eq_1 - star_of_number_less star_of_number_le star_of_number_eq - star_of_less_number star_of_le_number star_of_eq_number subsection {* Ordering and lattice classes *} @@ -984,9 +950,45 @@ subsection {* Number classes *} +instance star :: (numeral) numeral .. + +lemma star_numeral_def [transfer_unfold]: + "numeral k = star_of (numeral k)" +by (induct k, simp_all only: numeral.simps star_of_one star_of_add) + +lemma Standard_numeral [simp]: "numeral k \ Standard" +by (simp add: star_numeral_def) + +lemma star_of_numeral [simp]: "star_of (numeral k) = numeral k" +by transfer (rule refl) + +lemma star_neg_numeral_def [transfer_unfold]: + "neg_numeral k = star_of (neg_numeral k)" +by (simp only: neg_numeral_def star_of_minus star_of_numeral) + +lemma Standard_neg_numeral [simp]: "neg_numeral k \ Standard" +by (simp add: star_neg_numeral_def) + +lemma star_of_neg_numeral [simp]: "star_of (neg_numeral k) = neg_numeral k" +by transfer (rule refl) + lemma star_of_nat_def [transfer_unfold]: "of_nat n = star_of (of_nat n)" by (induct n, simp_all) +lemmas star_of_compare_numeral [simp] = + star_of_less [of "numeral k", simplified star_of_numeral] + star_of_le [of "numeral k", simplified star_of_numeral] + star_of_eq [of "numeral k", simplified star_of_numeral] + star_of_less [of _ "numeral k", simplified star_of_numeral] + star_of_le [of _ "numeral k", simplified star_of_numeral] + star_of_eq [of _ "numeral k", simplified star_of_numeral] + star_of_less [of "neg_numeral k", simplified star_of_numeral] + star_of_le [of "neg_numeral k", simplified star_of_numeral] + star_of_eq [of "neg_numeral k", simplified star_of_numeral] + star_of_less [of _ "neg_numeral k", simplified star_of_numeral] + star_of_le [of _ "neg_numeral k", simplified star_of_numeral] + star_of_eq [of _ "neg_numeral k", simplified star_of_numeral] for k + lemma Standard_of_nat [simp]: "of_nat n \ Standard" by (simp add: star_of_nat_def) @@ -1010,11 +1012,6 @@ instance star :: (ring_char_0) ring_char_0 .. -instance star :: (number_semiring) number_semiring -by (intro_classes, simp only: star_number_def star_of_nat_def number_of_int) - -instance star :: (number_ring) number_ring -by (intro_classes, simp only: star_number_def star_of_int_def number_of_eq) subsection {* Finite class *} diff -r 29e92b644d6c -r f067afe98049 src/HOL/Nat.thy --- a/src/HOL/Nat.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Nat.thy Mon Mar 26 10:56:56 2012 +0200 @@ -181,7 +181,7 @@ begin definition - One_nat_def [simp, code_post]: "1 = Suc 0" + One_nat_def [simp]: "1 = Suc 0" primrec times_nat where mult_0: "0 * n = (0\nat)" @@ -1782,4 +1782,6 @@ code_modulename Haskell Nat Arith +hide_const (open) of_nat_aux + end diff -r 29e92b644d6c -r f067afe98049 src/HOL/Nat_Numeral.thy --- a/src/HOL/Nat_Numeral.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Nat_Numeral.thy Mon Mar 26 10:56:56 2012 +0200 @@ -15,31 +15,13 @@ Arithmetic for naturals is reduced to that for the non-negative integers. *} -instantiation nat :: number_semiring -begin - -definition - nat_number_of_def [code_unfold, code del]: "number_of v = nat (number_of v)" - -instance proof - fix n show "number_of (int n) = (of_nat n :: nat)" - unfolding nat_number_of_def number_of_eq by simp -qed - -end - -lemma [code_post]: - "nat (number_of v) = number_of v" - unfolding nat_number_of_def .. - - subsection {* Special case: squares and cubes *} lemma numeral_2_eq_2: "2 = Suc (Suc 0)" - by (simp add: nat_number_of_def) + by (simp add: nat_number(2-4)) lemma numeral_3_eq_3: "3 = Suc (Suc (Suc 0))" - by (simp add: nat_number_of_def) + by (simp add: nat_number(2-4)) context power begin @@ -93,26 +75,21 @@ "(- a)\ = a\" by (simp add: power2_eq_square) -text{* - We cannot prove general results about the numeral @{term "-1"}, - so we have to use @{term "- 1"} instead. -*} - lemma power_minus1_even [simp]: - "(- 1) ^ (2*n) = 1" + "-1 ^ (2*n) = 1" proof (induct n) case 0 show ?case by simp next - case (Suc n) then show ?case by (simp add: power_add) + case (Suc n) then show ?case by (simp add: power_add power2_eq_square) qed lemma power_minus1_odd: - "(- 1) ^ Suc (2*n) = - 1" + "-1 ^ Suc (2*n) = -1" by simp lemma power_minus_even [simp]: "(-a) ^ (2*n) = a ^ (2*n)" - by (simp add: power_minus [of a]) + by (simp add: power_minus [of a]) end @@ -261,100 +238,31 @@ end lemma power2_sum: - fixes x y :: "'a::number_semiring" + fixes x y :: "'a::comm_semiring_1" shows "(x + y)\ = x\ + y\ + 2 * x * y" - by (simp add: algebra_simps power2_eq_square semiring_mult_2_right) + by (simp add: algebra_simps power2_eq_square mult_2_right) lemma power2_diff: - fixes x y :: "'a::number_ring" + fixes x y :: "'a::comm_ring_1" shows "(x - y)\ = x\ + y\ - 2 * x * y" by (simp add: ring_distribs power2_eq_square mult_2) (rule mult_commute) -subsection {* Predicate for negative binary numbers *} - -definition neg :: "int \ bool" where - "neg Z \ Z < 0" - -lemma not_neg_int [simp]: "~ neg (of_nat n)" -by (simp add: neg_def) - -lemma neg_zminus_int [simp]: "neg (- (of_nat (Suc n)))" -by (simp add: neg_def del: of_nat_Suc) - -lemmas neg_eq_less_0 = neg_def - -lemma not_neg_eq_ge_0: "(~neg x) = (0 \ x)" -by (simp add: neg_def linorder_not_less) - -text{*To simplify inequalities when Numeral1 can get simplified to 1*} - -lemma not_neg_0: "~ neg 0" -by (simp add: One_int_def neg_def) - -lemma not_neg_1: "~ neg 1" -by (simp add: neg_def linorder_not_less) - -lemma neg_nat: "neg z ==> nat z = 0" -by (simp add: neg_def order_less_imp_le) - -lemma not_neg_nat: "~ neg z ==> of_nat (nat z) = z" -by (simp add: linorder_not_less neg_def) - -text {* - If @{term Numeral0} is rewritten to 0 then this rule can't be applied: - @{term Numeral0} IS @{term "number_of Pls"} -*} - -lemma not_neg_number_of_Pls: "~ neg (number_of Int.Pls)" - by (simp add: neg_def) - -lemma neg_number_of_Min: "neg (number_of Int.Min)" - by (simp add: neg_def) - -lemma neg_number_of_Bit0: - "neg (number_of (Int.Bit0 w)) = neg (number_of w)" - by (simp add: neg_def) - -lemma neg_number_of_Bit1: - "neg (number_of (Int.Bit1 w)) = neg (number_of w)" - by (simp add: neg_def) - -lemmas neg_simps [simp] = - not_neg_0 not_neg_1 - not_neg_number_of_Pls neg_number_of_Min - neg_number_of_Bit0 neg_number_of_Bit1 - - subsection{*Function @{term nat}: Coercion from Type @{typ int} to @{typ nat}*} declare nat_1 [simp] -lemma nat_number_of [simp]: "nat (number_of w) = number_of w" - by (simp add: nat_number_of_def) - -lemma nat_numeral_0_eq_0: "Numeral0 = (0::nat)" (* FIXME delete candidate *) - by (fact semiring_numeral_0_eq_0) - -lemma nat_numeral_1_eq_1: "Numeral1 = (1::nat)" (* FIXME delete candidate *) - by (fact semiring_numeral_1_eq_1) - -lemma Numeral1_eq1_nat: - "(1::nat) = Numeral1" +lemma nat_neg_numeral [simp]: "nat (neg_numeral w) = 0" by simp lemma numeral_1_eq_Suc_0: "Numeral1 = Suc 0" - by (simp only: nat_numeral_1_eq_1 One_nat_def) + by simp subsection{*Function @{term int}: Coercion from Type @{typ nat} to @{typ int}*} -lemma int_nat_number_of [simp]: - "int (number_of v) = - (if neg (number_of v :: int) then 0 - else (number_of v :: int))" - unfolding nat_number_of_def number_of_is_id neg_def - by simp (* FIXME: redundant with of_nat_number_of_eq *) +lemma int_numeral: "int (numeral v) = numeral v" + by (rule of_nat_numeral) (* already simp *) lemma nonneg_int_cases: fixes k :: int assumes "0 \ k" obtains n where "k = of_nat n" @@ -368,149 +276,51 @@ done lemma Suc_nat_number_of_add: - "Suc (number_of v + n) = - (if neg (number_of v :: int) then 1+n else number_of (Int.succ v) + n)" - unfolding nat_number_of_def number_of_is_id neg_def numeral_simps - by (simp add: Suc_nat_eq_nat_zadd1 add_ac) - -lemma Suc_nat_number_of [simp]: - "Suc (number_of v) = - (if neg (number_of v :: int) then 1 else number_of (Int.succ v))" -apply (cut_tac n = 0 in Suc_nat_number_of_add) -apply (simp cong del: if_weak_cong) -done - - -subsubsection{*Addition *} - -lemma add_nat_number_of [simp]: - "(number_of v :: nat) + number_of v' = - (if v < Int.Pls then number_of v' - else if v' < Int.Pls then number_of v - else number_of (v + v'))" - unfolding nat_number_of_def number_of_is_id numeral_simps - by (simp add: nat_add_distrib) - -lemma nat_number_of_add_1 [simp]: - "number_of v + (1::nat) = - (if v < Int.Pls then 1 else number_of (Int.succ v))" - unfolding nat_number_of_def number_of_is_id numeral_simps - by (simp add: nat_add_distrib) + "Suc (numeral v + n) = numeral (v + Num.One) + n" + by simp -lemma nat_1_add_number_of [simp]: - "(1::nat) + number_of v = - (if v < Int.Pls then 1 else number_of (Int.succ v))" - unfolding nat_number_of_def number_of_is_id numeral_simps - by (simp add: nat_add_distrib) - -lemma nat_1_add_1 [simp]: "1 + 1 = (2::nat)" - by (rule semiring_one_add_one_is_two) - -text {* TODO: replace simp rules above with these generic ones: *} - -lemma semiring_add_number_of: - "\Int.Pls \ v; Int.Pls \ v'\ \ - (number_of v :: 'a::number_semiring) + number_of v' = number_of (v + v')" - unfolding Int.Pls_def - by (elim nonneg_int_cases, - simp only: number_of_int of_nat_add [symmetric]) - -lemma semiring_number_of_add_1: - "Int.Pls \ v \ - number_of v + (1::'a::number_semiring) = number_of (Int.succ v)" - unfolding Int.Pls_def Int.succ_def - by (elim nonneg_int_cases, - simp only: number_of_int add_commute [where b=1] of_nat_Suc [symmetric]) - -lemma semiring_1_add_number_of: - "Int.Pls \ v \ - (1::'a::number_semiring) + number_of v = number_of (Int.succ v)" - unfolding Int.Pls_def Int.succ_def - by (elim nonneg_int_cases, - simp only: number_of_int add_commute [where b=1] of_nat_Suc [symmetric]) +lemma Suc_numeral [simp]: + "Suc (numeral v) = numeral (v + Num.One)" + by simp subsubsection{*Subtraction *} lemma diff_nat_eq_if: "nat z - nat z' = - (if neg z' then nat z + (if z' < 0 then nat z else let d = z-z' in - if neg d then 0 else nat d)" -by (simp add: Let_def nat_diff_distrib [symmetric] neg_eq_less_0 not_neg_eq_ge_0) - - -lemma diff_nat_number_of [simp]: - "(number_of v :: nat) - number_of v' = - (if v' < Int.Pls then number_of v - else let d = number_of (v + uminus v') in - if neg d then 0 else nat d)" - unfolding nat_number_of_def number_of_is_id numeral_simps neg_def - by auto + if d < 0 then 0 else nat d)" +by (simp add: Let_def nat_diff_distrib [symmetric]) -lemma nat_number_of_diff_1 [simp]: - "number_of v - (1::nat) = - (if v \ Int.Pls then 0 else number_of (Int.pred v))" - unfolding nat_number_of_def number_of_is_id numeral_simps - by auto - - -subsubsection{*Multiplication *} +(* Int.nat_diff_distrib has too-strong premises *) +lemma nat_diff_distrib': "\0 \ x; 0 \ y\ \ nat (x - y) = nat x - nat y" +apply (rule int_int_eq [THEN iffD1], clarsimp) +apply (subst zdiff_int [symmetric]) +apply (rule nat_mono, simp_all) +done -lemma mult_nat_number_of [simp]: - "(number_of v :: nat) * number_of v' = - (if v < Int.Pls then 0 else number_of (v * v'))" - unfolding nat_number_of_def number_of_is_id numeral_simps - by (simp add: nat_mult_distrib) +lemma diff_nat_numeral [simp]: + "(numeral v :: nat) - numeral v' = nat (numeral v - numeral v')" + by (simp only: nat_diff_distrib' zero_le_numeral nat_numeral) -(* TODO: replace mult_nat_number_of with this next rule *) -lemma semiring_mult_number_of: - "\Int.Pls \ v; Int.Pls \ v'\ \ - (number_of v :: 'a::number_semiring) * number_of v' = number_of (v * v')" - unfolding Int.Pls_def - by (elim nonneg_int_cases, - simp only: number_of_int of_nat_mult [symmetric]) +lemma nat_numeral_diff_1 [simp]: + "numeral v - (1::nat) = nat (numeral v - 1)" + using diff_nat_numeral [of v Num.One] by simp subsection{*Comparisons*} -subsubsection{*Equals (=) *} - -lemma eq_nat_number_of [simp]: - "((number_of v :: nat) = number_of v') = - (if neg (number_of v :: int) then (number_of v' :: int) \ 0 - else if neg (number_of v' :: int) then (number_of v :: int) = 0 - else v = v')" - unfolding nat_number_of_def number_of_is_id neg_def - by auto - - -subsubsection{*Less-than (<) *} - -lemma less_nat_number_of [simp]: - "(number_of v :: nat) < number_of v' \ - (if v < v' then Int.Pls < v' else False)" - unfolding nat_number_of_def number_of_is_id numeral_simps - by auto - - -subsubsection{*Less-than-or-equal *} - -lemma le_nat_number_of [simp]: - "(number_of v :: nat) \ number_of v' \ - (if v \ v' then True else v \ Int.Pls)" - unfolding nat_number_of_def number_of_is_id numeral_simps - by auto - -(*Maps #n to n for n = 0, 1, 2*) -lemmas numerals = nat_numeral_0_eq_0 nat_numeral_1_eq_1 numeral_2_eq_2 +(*Maps #n to n for n = 1, 2*) +lemmas numerals = numeral_1_eq_1 [where 'a=nat] numeral_2_eq_2 subsection{*Powers with Numeric Exponents*} text{*Squares of literal numerals will be evaluated.*} -lemmas power2_eq_square_number_of [simp] = - power2_eq_square [of "number_of w"] for w +(* FIXME: replace with more general rules for powers of numerals *) +lemmas power2_eq_square_numeral [simp] = + power2_eq_square [of "numeral w"] for w text{*Simprules for comparisons where common factors can be cancelled.*} @@ -528,8 +338,8 @@ by simp (*Expresses a natural number constant as the Suc of another one. - NOT suitable for rewriting because n recurs in the condition.*) -lemmas expand_Suc = Suc_pred' [of "number_of v"] for v + NOT suitable for rewriting because n recurs on the right-hand side.*) +lemmas expand_Suc = Suc_pred' [of "numeral v", OF zero_less_numeral] for v subsubsection{*Arith *} @@ -539,7 +349,7 @@ lemma Suc_eq_plus1_left: "Suc n = 1 + n" unfolding One_nat_def by simp -(* These two can be useful when m = number_of... *) +(* These two can be useful when m = numeral... *) lemma add_eq_if: "(m::nat) + n = (if m=0 then n else Suc ((m - 1) + n))" unfolding One_nat_def by (cases m) simp_all @@ -551,231 +361,108 @@ unfolding One_nat_def by (cases m) simp_all -subsection{*Comparisons involving (0::nat) *} - -text{*Simplification already does @{term "n<0"}, @{term "n\0"} and @{term "0\n"}.*} - -lemma eq_number_of_0 [simp]: - "number_of v = (0::nat) \ v \ Int.Pls" - unfolding nat_number_of_def number_of_is_id numeral_simps - by auto - -lemma eq_0_number_of [simp]: - "(0::nat) = number_of v \ v \ Int.Pls" -by (rule trans [OF eq_sym_conv eq_number_of_0]) - -lemma less_0_number_of [simp]: - "(0::nat) < number_of v \ Int.Pls < v" - unfolding nat_number_of_def number_of_is_id numeral_simps - by simp - -lemma neg_imp_number_of_eq_0: "neg (number_of v :: int) ==> number_of v = (0::nat)" - by (simp del: semiring_numeral_0_eq_0 add: nat_numeral_0_eq_0 [symmetric]) - - subsection{*Comparisons involving @{term Suc} *} -lemma eq_number_of_Suc [simp]: - "(number_of v = Suc n) = - (let pv = number_of (Int.pred v) in - if neg pv then False else nat pv = n)" -apply (simp only: simp_thms Let_def neg_eq_less_0 linorder_not_less - number_of_pred nat_number_of_def - split add: split_if) -apply (rule_tac x = "number_of v" in spec) -apply (auto simp add: nat_eq_iff) -done +lemma eq_numeral_Suc [simp]: "numeral v = Suc n \ nat (numeral v - 1) = n" + by (subst expand_Suc, simp only: nat.inject nat_numeral_diff_1) -lemma Suc_eq_number_of [simp]: - "(Suc n = number_of v) = - (let pv = number_of (Int.pred v) in - if neg pv then False else nat pv = n)" -by (rule trans [OF eq_sym_conv eq_number_of_Suc]) +lemma Suc_eq_numeral [simp]: "Suc n = numeral v \ n = nat (numeral v - 1)" + by (subst expand_Suc, simp only: nat.inject nat_numeral_diff_1) -lemma less_number_of_Suc [simp]: - "(number_of v < Suc n) = - (let pv = number_of (Int.pred v) in - if neg pv then True else nat pv < n)" -apply (simp only: simp_thms Let_def neg_eq_less_0 linorder_not_less - number_of_pred nat_number_of_def - split add: split_if) -apply (rule_tac x = "number_of v" in spec) -apply (auto simp add: nat_less_iff) -done +lemma less_numeral_Suc [simp]: "numeral v < Suc n \ nat (numeral v - 1) < n" + by (subst expand_Suc, simp only: Suc_less_eq nat_numeral_diff_1) -lemma less_Suc_number_of [simp]: - "(Suc n < number_of v) = - (let pv = number_of (Int.pred v) in - if neg pv then False else n < nat pv)" -apply (simp only: simp_thms Let_def neg_eq_less_0 linorder_not_less - number_of_pred nat_number_of_def - split add: split_if) -apply (rule_tac x = "number_of v" in spec) -apply (auto simp add: zless_nat_eq_int_zless) -done +lemma less_Suc_numeral [simp]: "Suc n < numeral v \ n < nat (numeral v - 1)" + by (subst expand_Suc, simp only: Suc_less_eq nat_numeral_diff_1) -lemma le_number_of_Suc [simp]: - "(number_of v <= Suc n) = - (let pv = number_of (Int.pred v) in - if neg pv then True else nat pv <= n)" -by (simp add: Let_def linorder_not_less [symmetric]) +lemma le_numeral_Suc [simp]: "numeral v \ Suc n \ nat (numeral v - 1) \ n" + by (subst expand_Suc, simp only: Suc_le_mono nat_numeral_diff_1) -lemma le_Suc_number_of [simp]: - "(Suc n <= number_of v) = - (let pv = number_of (Int.pred v) in - if neg pv then False else n <= nat pv)" -by (simp add: Let_def linorder_not_less [symmetric]) - - -lemma eq_number_of_Pls_Min: "(Numeral0 ::int) ~= number_of Int.Min" -by auto - +lemma le_Suc_numeral [simp]: "Suc n \ numeral v \ n \ nat (numeral v - 1)" + by (subst expand_Suc, simp only: Suc_le_mono nat_numeral_diff_1) subsection{*Max and Min Combined with @{term Suc} *} -lemma max_number_of_Suc [simp]: - "max (Suc n) (number_of v) = - (let pv = number_of (Int.pred v) in - if neg pv then Suc n else Suc(max n (nat pv)))" -apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def - split add: split_if nat.split) -apply (rule_tac x = "number_of v" in spec) -apply auto -done - -lemma max_Suc_number_of [simp]: - "max (number_of v) (Suc n) = - (let pv = number_of (Int.pred v) in - if neg pv then Suc n else Suc(max (nat pv) n))" -apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def - split add: split_if nat.split) -apply (rule_tac x = "number_of v" in spec) -apply auto -done - -lemma min_number_of_Suc [simp]: - "min (Suc n) (number_of v) = - (let pv = number_of (Int.pred v) in - if neg pv then 0 else Suc(min n (nat pv)))" -apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def - split add: split_if nat.split) -apply (rule_tac x = "number_of v" in spec) -apply auto -done - -lemma min_Suc_number_of [simp]: - "min (number_of v) (Suc n) = - (let pv = number_of (Int.pred v) in - if neg pv then 0 else Suc(min (nat pv) n))" -apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def - split add: split_if nat.split) -apply (rule_tac x = "number_of v" in spec) -apply auto -done +lemma max_Suc_numeral [simp]: + "max (Suc n) (numeral v) = Suc (max n (nat (numeral v - 1)))" + by (subst expand_Suc, simp only: max_Suc_Suc nat_numeral_diff_1) + +lemma max_numeral_Suc [simp]: + "max (numeral v) (Suc n) = Suc (max (nat (numeral v - 1)) n)" + by (subst expand_Suc, simp only: max_Suc_Suc nat_numeral_diff_1) + +lemma min_Suc_numeral [simp]: + "min (Suc n) (numeral v) = Suc (min n (nat (numeral v - 1)))" + by (subst expand_Suc, simp only: min_Suc_Suc nat_numeral_diff_1) + +lemma min_numeral_Suc [simp]: + "min (numeral v) (Suc n) = Suc (min (nat (numeral v - 1)) n)" + by (subst expand_Suc, simp only: min_Suc_Suc nat_numeral_diff_1) subsection{*Literal arithmetic involving powers*} -lemma power_nat_number_of: - "(number_of v :: nat) ^ n = - (if neg (number_of v :: int) then 0^n else nat ((number_of v :: int) ^ n))" -by (simp only: simp_thms neg_nat not_neg_eq_ge_0 nat_number_of_def nat_power_eq - split add: split_if cong: imp_cong) +(* TODO: replace with more generic rule for powers of numerals *) +lemma power_nat_numeral: + "(numeral v :: nat) ^ n = nat ((numeral v :: int) ^ n)" + by (simp only: nat_power_eq zero_le_numeral nat_numeral) - -lemmas power_nat_number_of_number_of = power_nat_number_of [of _ "number_of w"] for w -declare power_nat_number_of_number_of [simp] - +lemmas power_nat_numeral_numeral = power_nat_numeral [of _ "numeral w"] for w +declare power_nat_numeral_numeral [simp] text{*For arbitrary rings*} -lemma power_number_of_even: +lemma power_numeral_even: fixes z :: "'a::monoid_mult" - shows "z ^ number_of (Int.Bit0 w) = (let w = z ^ (number_of w) in w * w)" -by (cases "w \ 0") (auto simp add: Let_def Bit0_def nat_number_of_def number_of_is_id - nat_add_distrib power_add simp del: nat_number_of) + shows "z ^ numeral (Num.Bit0 w) = (let w = z ^ (numeral w) in w * w)" + unfolding numeral_Bit0 power_add Let_def .. -lemma power_number_of_odd: +lemma power_numeral_odd: fixes z :: "'a::monoid_mult" - shows "z ^ number_of (Int.Bit1 w) = (if (0::int) <= number_of w - then (let w = z ^ (number_of w) in z * w * w) else 1)" -unfolding Let_def Bit1_def nat_number_of_def number_of_is_id -apply (cases "0 <= w") -apply (simp only: mult_assoc nat_add_distrib power_add, simp) -apply (simp add: not_le mult_2 [symmetric] add_assoc) -done + shows "z ^ numeral (Num.Bit1 w) = (let w = z ^ (numeral w) in z * w * w)" + unfolding numeral_Bit1 One_nat_def add_Suc_right add_0_right + unfolding power_Suc power_add Let_def mult_assoc .. -lemmas zpower_number_of_even = power_number_of_even [where 'a=int] -lemmas zpower_number_of_odd = power_number_of_odd [where 'a=int] - -lemmas power_number_of_even_number_of [simp] = - power_number_of_even [of "number_of v"] for v +lemmas zpower_numeral_even = power_numeral_even [where 'a=int] +lemmas zpower_numeral_odd = power_numeral_odd [where 'a=int] -lemmas power_number_of_odd_number_of [simp] = - power_number_of_odd [of "number_of v"] for v +lemmas power_numeral_even_numeral [simp] = + power_numeral_even [of "numeral v"] for v -lemma nat_number_of_Pls: "Numeral0 = (0::nat)" - by (simp add: nat_number_of_def) - -lemma nat_number_of_Min [no_atp]: "number_of Int.Min = (0::nat)" - apply (simp only: number_of_Min nat_number_of_def nat_zminus_int) - done +lemmas power_numeral_odd_numeral [simp] = + power_numeral_odd [of "numeral v"] for v -lemma nat_number_of_Bit0: - "number_of (Int.Bit0 w) = (let n::nat = number_of w in n + n)" -by (cases "w \ 0") (auto simp add: Let_def Bit0_def nat_number_of_def number_of_is_id - nat_add_distrib simp del: nat_number_of) +lemma nat_numeral_Bit0: + "numeral (Num.Bit0 w) = (let n::nat = numeral w in n + n)" + unfolding numeral_Bit0 Let_def .. -lemma nat_number_of_Bit1: - "number_of (Int.Bit1 w) = - (if neg (number_of w :: int) then 0 - else let n = number_of w in Suc (n + n))" -unfolding Let_def Bit1_def nat_number_of_def number_of_is_id neg_def -apply (cases "w < 0") -apply (simp add: mult_2 [symmetric] add_assoc) -apply (simp only: nat_add_distrib, simp) -done +lemma nat_numeral_Bit1: + "numeral (Num.Bit1 w) = (let n = numeral w in Suc (n + n))" + unfolding numeral_Bit1 Let_def by simp lemmas eval_nat_numeral = - nat_number_of_Bit0 nat_number_of_Bit1 + nat_numeral_Bit0 nat_numeral_Bit1 lemmas nat_arith = - add_nat_number_of - diff_nat_number_of - mult_nat_number_of - eq_nat_number_of - less_nat_number_of + diff_nat_numeral lemmas semiring_norm = - Let_def arith_simps nat_arith rel_simps neg_simps if_False - if_True add_0 add_Suc add_number_of_left mult_number_of_left + Let_def arith_simps nat_arith rel_simps + if_False if_True + add_0 add_Suc add_numeral_left + add_neg_numeral_left mult_numeral_left numeral_1_eq_1 [symmetric] Suc_eq_plus1 - numeral_0_eq_0 [symmetric] numerals [symmetric] - not_iszero_Numeral1 + eq_numeral_iff_iszero not_iszero_Numeral1 lemma Let_Suc [simp]: "Let (Suc n) f == f (Suc n)" by (fact Let_def) -lemma power_m1_even: "(-1) ^ (2*n) = (1::'a::{number_ring})" - by (simp only: number_of_Min power_minus1_even) - -lemma power_m1_odd: "(-1) ^ Suc(2*n) = (-1::'a::{number_ring})" - by (simp only: number_of_Min power_minus1_odd) +lemma power_m1_even: "(-1) ^ (2*n) = (1::'a::ring_1)" + by (fact power_minus1_even) (* FIXME: duplicate *) -lemma nat_number_of_add_left: - "number_of v + (number_of v' + (k::nat)) = - (if neg (number_of v :: int) then number_of v' + k - else if neg (number_of v' :: int) then number_of v + k - else number_of (v + v') + k)" -by (auto simp add: neg_def) - -lemma nat_number_of_mult_left: - "number_of v * (number_of v' * (k::nat)) = - (if v < Int.Pls then 0 - else number_of (v * v') * k)" -by (auto simp add: not_less Pls_def nat_number_of_def number_of_is_id - nat_mult_distrib simp del: nat_number_of) +lemma power_m1_odd: "(-1) ^ Suc(2*n) = (-1::'a::ring_1)" + by (fact power_minus1_odd) (* FIXME: duplicate *) subsection{*Literal arithmetic and @{term of_nat}*} @@ -784,52 +471,18 @@ "0 \ x ==> of_nat (nat (2 * x)) = of_nat (nat x) + of_nat (nat x)" by (simp only: mult_2 nat_add_distrib of_nat_add) -lemma nat_numeral_m1_eq_0: "-1 = (0::nat)" -by (simp only: nat_number_of_def) - -lemma of_nat_number_of_lemma: - "of_nat (number_of v :: nat) = - (if 0 \ (number_of v :: int) - then (number_of v :: 'a :: number_semiring) - else 0)" - by (auto simp add: int_number_of_def nat_number_of_def number_of_int - elim!: nonneg_int_cases) - -lemma of_nat_number_of_eq [simp]: - "of_nat (number_of v :: nat) = - (if neg (number_of v :: int) then 0 - else (number_of v :: 'a :: number_semiring))" - by (simp only: of_nat_number_of_lemma neg_def, simp) - subsubsection{*For simplifying @{term "Suc m - K"} and @{term "K - Suc m"}*} text{*Where K above is a literal*} -lemma Suc_diff_eq_diff_pred: "Numeral0 < n ==> Suc m - n = m - (n - Numeral1)" +lemma Suc_diff_eq_diff_pred: "0 < n ==> Suc m - n = m - (n - Numeral1)" by (simp split: nat_diff_split) -text {*Now just instantiating @{text n} to @{text "number_of v"} does - the right simplification, but with some redundant inequality - tests.*} -lemma neg_number_of_pred_iff_0: - "neg (number_of (Int.pred v)::int) = (number_of v = (0::nat))" -apply (subgoal_tac "neg (number_of (Int.pred v)) = (number_of v < Suc 0) ") -apply (simp only: less_Suc_eq_le le_0_eq) -apply (subst less_number_of_Suc, simp) -done - text{*No longer required as a simprule because of the @{text inverse_fold} simproc*} -lemma Suc_diff_number_of: - "Int.Pls < v ==> - Suc m - (number_of v) = m - (number_of (Int.pred v))" -apply (subst Suc_diff_eq_diff_pred) -apply simp -apply (simp del: semiring_numeral_1_eq_1) -apply (auto simp only: diff_nat_number_of less_0_number_of [symmetric] - neg_number_of_pred_iff_0) -done +lemma Suc_diff_numeral: "Suc m - (numeral v) = m - (numeral v - 1)" + by (subst expand_Suc, simp only: diff_Suc_Suc) lemma diff_Suc_eq_diff_pred: "m - Suc n = (m - 1) - n" by (simp split: nat_diff_split) @@ -837,45 +490,22 @@ subsubsection{*For @{term nat_case} and @{term nat_rec}*} -lemma nat_case_number_of [simp]: - "nat_case a f (number_of v) = - (let pv = number_of (Int.pred v) in - if neg pv then a else f (nat pv))" -by (simp split add: nat.split add: Let_def neg_number_of_pred_iff_0) +lemma nat_case_numeral [simp]: + "nat_case a f (numeral v) = (let pv = nat (numeral v - 1) in f pv)" + by (subst expand_Suc, simp only: nat.cases nat_numeral_diff_1 Let_def) lemma nat_case_add_eq_if [simp]: - "nat_case a f ((number_of v) + n) = - (let pv = number_of (Int.pred v) in - if neg pv then nat_case a f n else f (nat pv + n))" -apply (subst add_eq_if) -apply (simp split add: nat.split - del: semiring_numeral_1_eq_1 - add: semiring_numeral_1_eq_1 [symmetric] - numeral_1_eq_Suc_0 [symmetric] - neg_number_of_pred_iff_0) -done + "nat_case a f ((numeral v) + n) = (let pv = nat (numeral v - 1) in f (pv + n))" + by (subst expand_Suc, simp only: nat.cases nat_numeral_diff_1 Let_def add_Suc) -lemma nat_rec_number_of [simp]: - "nat_rec a f (number_of v) = - (let pv = number_of (Int.pred v) in - if neg pv then a else f (nat pv) (nat_rec a f (nat pv)))" -apply (case_tac " (number_of v) ::nat") -apply (simp_all (no_asm_simp) add: Let_def neg_number_of_pred_iff_0) -apply (simp split add: split_if_asm) -done +lemma nat_rec_numeral [simp]: + "nat_rec a f (numeral v) = (let pv = nat (numeral v - 1) in f pv (nat_rec a f pv))" + by (subst expand_Suc, simp only: nat_rec_Suc nat_numeral_diff_1 Let_def) lemma nat_rec_add_eq_if [simp]: - "nat_rec a f (number_of v + n) = - (let pv = number_of (Int.pred v) in - if neg pv then nat_rec a f n - else f (nat pv + n) (nat_rec a f (nat pv + n)))" -apply (subst add_eq_if) -apply (simp split add: nat.split - del: semiring_numeral_1_eq_1 - add: semiring_numeral_1_eq_1 [symmetric] - numeral_1_eq_Suc_0 [symmetric] - neg_number_of_pred_iff_0) -done + "nat_rec a f (numeral v + n) = + (let pv = nat (numeral v - 1) in f (pv + n) (nat_rec a f (pv + n)))" + by (subst expand_Suc, simp only: nat_rec_Suc nat_numeral_diff_1 Let_def add_Suc) subsubsection{*Various Other Lemmas*} @@ -887,14 +517,14 @@ text{*Lemmas for specialist use, NOT as default simprules*} lemma nat_mult_2: "2 * z = (z+z::nat)" -by (rule semiring_mult_2) +by (rule mult_2) (* FIXME: duplicate *) lemma nat_mult_2_right: "z * 2 = (z+z::nat)" -by (rule semiring_mult_2_right) +by (rule mult_2_right) (* FIXME: duplicate *) text{*Case analysis on @{term "n<2"}*} lemma less_2_cases: "(n::nat) < 2 ==> n = 0 | n = Suc 0" -by (auto simp add: nat_1_add_1 [symmetric]) +by (auto simp add: numeral_2_eq_2) text{*Removal of Small Numerals: 0, 1 and (in additive positions) 2*} @@ -908,4 +538,8 @@ lemma Suc3_eq_add_3: "Suc (Suc (Suc n)) = 3 + n" by simp +text{*Legacy theorems*} + +lemmas nat_1_add_1 = one_add_one [where 'a=nat] + end diff -r 29e92b644d6c -r f067afe98049 src/HOL/Nitpick_Examples/Integer_Nits.thy diff -r 29e92b644d6c -r f067afe98049 src/HOL/Nitpick_Examples/Manual_Nits.thy --- a/src/HOL/Nitpick_Examples/Manual_Nits.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Nitpick_Examples/Manual_Nits.thy Mon Mar 26 10:56:56 2012 +0200 @@ -115,6 +115,7 @@ "add_raw \ \(x, y) (u, v). (x + (u\nat), y + (v\nat))" quotient_definition "add\my_int \ my_int \ my_int" is add_raw +unfolding add_raw_def by auto lemma "add x y = add x x" nitpick [show_datatypes, expect = genuine] diff -r 29e92b644d6c -r f067afe98049 src/HOL/Nominal/Nominal.thy --- a/src/HOL/Nominal/Nominal.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Nominal/Nominal.thy Mon Mar 26 10:56:56 2012 +0200 @@ -3481,7 +3481,7 @@ by (auto simp add: perm_nat_def) lemma numeral_nat_eqvt: - shows "pi\((number_of n)::nat) = number_of n" + shows "pi\((numeral n)::nat) = numeral n" by (simp add: perm_nat_def perm_int_def) lemma max_nat_eqvt: @@ -3523,7 +3523,11 @@ by (simp add: perm_int_def) lemma numeral_int_eqvt: - shows "pi\((number_of n)::int) = number_of n" + shows "pi\((numeral n)::int) = numeral n" +by (simp add: perm_int_def perm_int_def) + +lemma neg_numeral_int_eqvt: + shows "pi\((neg_numeral n)::int) = neg_numeral n" by (simp add: perm_int_def perm_int_def) lemma max_int_eqvt: @@ -3589,7 +3593,7 @@ (* the lemmas numeral_nat_eqvt numeral_int_eqvt do not conform with the *) (* usual form of an eqvt-lemma, but they are needed for analysing *) (* permutations on nats and ints *) -lemmas [eqvt_force] = numeral_nat_eqvt numeral_int_eqvt +lemmas [eqvt_force] = numeral_nat_eqvt numeral_int_eqvt neg_numeral_int_eqvt (***************************************) (* setup for the individial atom-kinds *) diff -r 29e92b644d6c -r f067afe98049 src/HOL/Num.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Num.thy Mon Mar 26 10:56:56 2012 +0200 @@ -0,0 +1,1021 @@ +(* Title: HOL/Num.thy + Author: Florian Haftmann + Author: Brian Huffman +*) + +header {* Binary Numerals *} + +theory Num +imports Datatype Power +begin + +subsection {* The @{text num} type *} + +datatype num = One | Bit0 num | Bit1 num + +text {* Increment function for type @{typ num} *} + +primrec inc :: "num \ num" where + "inc One = Bit0 One" | + "inc (Bit0 x) = Bit1 x" | + "inc (Bit1 x) = Bit0 (inc x)" + +text {* Converting between type @{typ num} and type @{typ nat} *} + +primrec nat_of_num :: "num \ nat" where + "nat_of_num One = Suc 0" | + "nat_of_num (Bit0 x) = nat_of_num x + nat_of_num x" | + "nat_of_num (Bit1 x) = Suc (nat_of_num x + nat_of_num x)" + +primrec num_of_nat :: "nat \ num" where + "num_of_nat 0 = One" | + "num_of_nat (Suc n) = (if 0 < n then inc (num_of_nat n) else One)" + +lemma nat_of_num_pos: "0 < nat_of_num x" + by (induct x) simp_all + +lemma nat_of_num_neq_0: " nat_of_num x \ 0" + by (induct x) simp_all + +lemma nat_of_num_inc: "nat_of_num (inc x) = Suc (nat_of_num x)" + by (induct x) simp_all + +lemma num_of_nat_double: + "0 < n \ num_of_nat (n + n) = Bit0 (num_of_nat n)" + by (induct n) simp_all + +text {* + Type @{typ num} is isomorphic to the strictly positive + natural numbers. +*} + +lemma nat_of_num_inverse: "num_of_nat (nat_of_num x) = x" + by (induct x) (simp_all add: num_of_nat_double nat_of_num_pos) + +lemma num_of_nat_inverse: "0 < n \ nat_of_num (num_of_nat n) = n" + by (induct n) (simp_all add: nat_of_num_inc) + +lemma num_eq_iff: "x = y \ nat_of_num x = nat_of_num y" + apply safe + apply (drule arg_cong [where f=num_of_nat]) + apply (simp add: nat_of_num_inverse) + done + +lemma num_induct [case_names One inc]: + fixes P :: "num \ bool" + assumes One: "P One" + and inc: "\x. P x \ P (inc x)" + shows "P x" +proof - + obtain n where n: "Suc n = nat_of_num x" + by (cases "nat_of_num x", simp_all add: nat_of_num_neq_0) + have "P (num_of_nat (Suc n))" + proof (induct n) + case 0 show ?case using One by simp + next + case (Suc n) + then have "P (inc (num_of_nat (Suc n)))" by (rule inc) + then show "P (num_of_nat (Suc (Suc n)))" by simp + qed + with n show "P x" + by (simp add: nat_of_num_inverse) +qed + +text {* + From now on, there are two possible models for @{typ num}: + as positive naturals (rule @{text "num_induct"}) + and as digit representation (rules @{text "num.induct"}, @{text "num.cases"}). +*} + + +subsection {* Numeral operations *} + +instantiation num :: "{plus,times,linorder}" +begin + +definition [code del]: + "m + n = num_of_nat (nat_of_num m + nat_of_num n)" + +definition [code del]: + "m * n = num_of_nat (nat_of_num m * nat_of_num n)" + +definition [code del]: + "m \ n \ nat_of_num m \ nat_of_num n" + +definition [code del]: + "m < n \ nat_of_num m < nat_of_num n" + +instance + by (default, auto simp add: less_num_def less_eq_num_def num_eq_iff) + +end + +lemma nat_of_num_add: "nat_of_num (x + y) = nat_of_num x + nat_of_num y" + unfolding plus_num_def + by (intro num_of_nat_inverse add_pos_pos nat_of_num_pos) + +lemma nat_of_num_mult: "nat_of_num (x * y) = nat_of_num x * nat_of_num y" + unfolding times_num_def + by (intro num_of_nat_inverse mult_pos_pos nat_of_num_pos) + +lemma add_num_simps [simp, code]: + "One + One = Bit0 One" + "One + Bit0 n = Bit1 n" + "One + Bit1 n = Bit0 (n + One)" + "Bit0 m + One = Bit1 m" + "Bit0 m + Bit0 n = Bit0 (m + n)" + "Bit0 m + Bit1 n = Bit1 (m + n)" + "Bit1 m + One = Bit0 (m + One)" + "Bit1 m + Bit0 n = Bit1 (m + n)" + "Bit1 m + Bit1 n = Bit0 (m + n + One)" + by (simp_all add: num_eq_iff nat_of_num_add) + +lemma mult_num_simps [simp, code]: + "m * One = m" + "One * n = n" + "Bit0 m * Bit0 n = Bit0 (Bit0 (m * n))" + "Bit0 m * Bit1 n = Bit0 (m * Bit1 n)" + "Bit1 m * Bit0 n = Bit0 (Bit1 m * n)" + "Bit1 m * Bit1 n = Bit1 (m + n + Bit0 (m * n))" + by (simp_all add: num_eq_iff nat_of_num_add + nat_of_num_mult left_distrib right_distrib) + +lemma eq_num_simps: + "One = One \ True" + "One = Bit0 n \ False" + "One = Bit1 n \ False" + "Bit0 m = One \ False" + "Bit1 m = One \ False" + "Bit0 m = Bit0 n \ m = n" + "Bit0 m = Bit1 n \ False" + "Bit1 m = Bit0 n \ False" + "Bit1 m = Bit1 n \ m = n" + by simp_all + +lemma le_num_simps [simp, code]: + "One \ n \ True" + "Bit0 m \ One \ False" + "Bit1 m \ One \ False" + "Bit0 m \ Bit0 n \ m \ n" + "Bit0 m \ Bit1 n \ m \ n" + "Bit1 m \ Bit1 n \ m \ n" + "Bit1 m \ Bit0 n \ m < n" + using nat_of_num_pos [of n] nat_of_num_pos [of m] + by (auto simp add: less_eq_num_def less_num_def) + +lemma less_num_simps [simp, code]: + "m < One \ False" + "One < Bit0 n \ True" + "One < Bit1 n \ True" + "Bit0 m < Bit0 n \ m < n" + "Bit0 m < Bit1 n \ m \ n" + "Bit1 m < Bit1 n \ m < n" + "Bit1 m < Bit0 n \ m < n" + using nat_of_num_pos [of n] nat_of_num_pos [of m] + by (auto simp add: less_eq_num_def less_num_def) + +text {* Rules using @{text One} and @{text inc} as constructors *} + +lemma add_One: "x + One = inc x" + by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc) + +lemma add_One_commute: "One + n = n + One" + by (induct n) simp_all + +lemma add_inc: "x + inc y = inc (x + y)" + by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc) + +lemma mult_inc: "x * inc y = x * y + x" + by (simp add: num_eq_iff nat_of_num_mult nat_of_num_add nat_of_num_inc) + +text {* The @{const num_of_nat} conversion *} + +lemma num_of_nat_One: + "n \ 1 \ num_of_nat n = Num.One" + by (cases n) simp_all + +lemma num_of_nat_plus_distrib: + "0 < m \ 0 < n \ num_of_nat (m + n) = num_of_nat m + num_of_nat n" + by (induct n) (auto simp add: add_One add_One_commute add_inc) + +text {* A double-and-decrement function *} + +primrec BitM :: "num \ num" where + "BitM One = One" | + "BitM (Bit0 n) = Bit1 (BitM n)" | + "BitM (Bit1 n) = Bit1 (Bit0 n)" + +lemma BitM_plus_one: "BitM n + One = Bit0 n" + by (induct n) simp_all + +lemma one_plus_BitM: "One + BitM n = Bit0 n" + unfolding add_One_commute BitM_plus_one .. + +text {* Squaring and exponentiation *} + +primrec sqr :: "num \ num" where + "sqr One = One" | + "sqr (Bit0 n) = Bit0 (Bit0 (sqr n))" | + "sqr (Bit1 n) = Bit1 (Bit0 (sqr n + n))" + +primrec pow :: "num \ num \ num" where + "pow x One = x" | + "pow x (Bit0 y) = sqr (pow x y)" | + "pow x (Bit1 y) = x * sqr (pow x y)" + +lemma nat_of_num_sqr: "nat_of_num (sqr x) = nat_of_num x * nat_of_num x" + by (induct x, simp_all add: algebra_simps nat_of_num_add) + +lemma sqr_conv_mult: "sqr x = x * x" + by (simp add: num_eq_iff nat_of_num_sqr nat_of_num_mult) + + +subsection {* Numary numerals *} + +text {* + We embed numary representations into a generic algebraic + structure using @{text numeral}. +*} + +class numeral = one + semigroup_add +begin + +primrec numeral :: "num \ 'a" where + numeral_One: "numeral One = 1" | + numeral_Bit0: "numeral (Bit0 n) = numeral n + numeral n" | + numeral_Bit1: "numeral (Bit1 n) = numeral n + numeral n + 1" + +lemma one_plus_numeral_commute: "1 + numeral x = numeral x + 1" + apply (induct x) + apply simp + apply (simp add: add_assoc [symmetric], simp add: add_assoc) + apply (simp add: add_assoc [symmetric], simp add: add_assoc) + done + +lemma numeral_inc: "numeral (inc x) = numeral x + 1" +proof (induct x) + case (Bit1 x) + have "numeral x + (1 + numeral x) + 1 = numeral x + (numeral x + 1) + 1" + by (simp only: one_plus_numeral_commute) + with Bit1 show ?case + by (simp add: add_assoc) +qed simp_all + +declare numeral.simps [simp del] + +abbreviation "Numeral1 \ numeral One" + +declare numeral_One [code_post] + +end + +text {* Negative numerals. *} + +class neg_numeral = numeral + group_add +begin + +definition neg_numeral :: "num \ 'a" where + "neg_numeral k = - numeral k" + +end + +text {* Numeral syntax. *} + +syntax + "_Numeral" :: "num_const \ 'a" ("_") + +parse_translation {* +let + fun num_of_int n = if n > 0 then case IntInf.quotRem (n, 2) + of (0, 1) => Syntax.const @{const_name One} + | (n, 0) => Syntax.const @{const_name Bit0} $ num_of_int n + | (n, 1) => Syntax.const @{const_name Bit1} $ num_of_int n + else raise Match; + val pos = Syntax.const @{const_name numeral} + val neg = Syntax.const @{const_name neg_numeral} + val one = Syntax.const @{const_name Groups.one} + val zero = Syntax.const @{const_name Groups.zero} + fun numeral_tr [(c as Const (@{syntax_const "_constrain"}, _)) $ t $ u] = + c $ numeral_tr [t] $ u + | numeral_tr [Const (num, _)] = + let + val {value, ...} = Lexicon.read_xnum num; + in + if value = 0 then zero else + if value > 0 + then pos $ num_of_int value + else neg $ num_of_int (~value) + end + | numeral_tr ts = raise TERM ("numeral_tr", ts); +in [("_Numeral", numeral_tr)] end +*} + +typed_print_translation (advanced) {* +let + fun dest_num (Const (@{const_syntax Bit0}, _) $ n) = 2 * dest_num n + | dest_num (Const (@{const_syntax Bit1}, _) $ n) = 2 * dest_num n + 1 + | dest_num (Const (@{const_syntax One}, _)) = 1; + fun num_tr' sign ctxt T [n] = + let + val k = dest_num n; + val t' = Syntax.const @{syntax_const "_Numeral"} $ + Syntax.free (sign ^ string_of_int k); + in + case T of + Type (@{type_name fun}, [_, T']) => + if not (Config.get ctxt show_types) andalso can Term.dest_Type T' then t' + else Syntax.const @{syntax_const "_constrain"} $ t' $ Syntax_Phases.term_of_typ ctxt T' + | T' => if T' = dummyT then t' else raise Match + end; +in [(@{const_syntax numeral}, num_tr' ""), + (@{const_syntax neg_numeral}, num_tr' "-")] end +*} + +subsection {* Class-specific numeral rules *} + +text {* + @{const numeral} is a morphism. +*} + +subsubsection {* Structures with addition: class @{text numeral} *} + +context numeral +begin + +lemma numeral_add: "numeral (m + n) = numeral m + numeral n" + by (induct n rule: num_induct) + (simp_all only: numeral_One add_One add_inc numeral_inc add_assoc) + +lemma numeral_plus_numeral: "numeral m + numeral n = numeral (m + n)" + by (rule numeral_add [symmetric]) + +lemma numeral_plus_one: "numeral n + 1 = numeral (n + One)" + using numeral_add [of n One] by (simp add: numeral_One) + +lemma one_plus_numeral: "1 + numeral n = numeral (One + n)" + using numeral_add [of One n] by (simp add: numeral_One) + +lemma one_add_one: "1 + 1 = 2" + using numeral_add [of One One] by (simp add: numeral_One) + +lemmas add_numeral_special = + numeral_plus_one one_plus_numeral one_add_one + +end + +subsubsection {* + Structures with negation: class @{text neg_numeral} +*} + +context neg_numeral +begin + +text {* Numerals form an abelian subgroup. *} + +inductive is_num :: "'a \ bool" where + "is_num 1" | + "is_num x \ is_num (- x)" | + "\is_num x; is_num y\ \ is_num (x + y)" + +lemma is_num_numeral: "is_num (numeral k)" + by (induct k, simp_all add: numeral.simps is_num.intros) + +lemma is_num_add_commute: + "\is_num x; is_num y\ \ x + y = y + x" + apply (induct x rule: is_num.induct) + apply (induct y rule: is_num.induct) + apply simp + apply (rule_tac a=x in add_left_imp_eq) + apply (rule_tac a=x in add_right_imp_eq) + apply (simp add: add_assoc minus_add_cancel) + apply (simp add: add_assoc [symmetric], simp add: add_assoc) + apply (rule_tac a=x in add_left_imp_eq) + apply (rule_tac a=x in add_right_imp_eq) + apply (simp add: add_assoc minus_add_cancel add_minus_cancel) + apply (simp add: add_assoc, simp add: add_assoc [symmetric]) + done + +lemma is_num_add_left_commute: + "\is_num x; is_num y\ \ x + (y + z) = y + (x + z)" + by (simp only: add_assoc [symmetric] is_num_add_commute) + +lemmas is_num_normalize = + add_assoc is_num_add_commute is_num_add_left_commute + is_num.intros is_num_numeral + diff_minus minus_add add_minus_cancel minus_add_cancel + +definition dbl :: "'a \ 'a" where "dbl x = x + x" +definition dbl_inc :: "'a \ 'a" where "dbl_inc x = x + x + 1" +definition dbl_dec :: "'a \ 'a" where "dbl_dec x = x + x - 1" + +definition sub :: "num \ num \ 'a" where + "sub k l = numeral k - numeral l" + +lemma numeral_BitM: "numeral (BitM n) = numeral (Bit0 n) - 1" + by (simp only: BitM_plus_one [symmetric] numeral_add numeral_One eq_diff_eq) + +lemma dbl_simps [simp]: + "dbl (neg_numeral k) = neg_numeral (Bit0 k)" + "dbl 0 = 0" + "dbl 1 = 2" + "dbl (numeral k) = numeral (Bit0 k)" + unfolding dbl_def neg_numeral_def numeral.simps + by (simp_all add: minus_add) + +lemma dbl_inc_simps [simp]: + "dbl_inc (neg_numeral k) = neg_numeral (BitM k)" + "dbl_inc 0 = 1" + "dbl_inc 1 = 3" + "dbl_inc (numeral k) = numeral (Bit1 k)" + unfolding dbl_inc_def neg_numeral_def numeral.simps numeral_BitM + by (simp_all add: is_num_normalize) + +lemma dbl_dec_simps [simp]: + "dbl_dec (neg_numeral k) = neg_numeral (Bit1 k)" + "dbl_dec 0 = -1" + "dbl_dec 1 = 1" + "dbl_dec (numeral k) = numeral (BitM k)" + unfolding dbl_dec_def neg_numeral_def numeral.simps numeral_BitM + by (simp_all add: is_num_normalize) + +lemma sub_num_simps [simp]: + "sub One One = 0" + "sub One (Bit0 l) = neg_numeral (BitM l)" + "sub One (Bit1 l) = neg_numeral (Bit0 l)" + "sub (Bit0 k) One = numeral (BitM k)" + "sub (Bit1 k) One = numeral (Bit0 k)" + "sub (Bit0 k) (Bit0 l) = dbl (sub k l)" + "sub (Bit0 k) (Bit1 l) = dbl_dec (sub k l)" + "sub (Bit1 k) (Bit0 l) = dbl_inc (sub k l)" + "sub (Bit1 k) (Bit1 l) = dbl (sub k l)" + unfolding dbl_def dbl_dec_def dbl_inc_def sub_def + unfolding neg_numeral_def numeral.simps numeral_BitM + by (simp_all add: is_num_normalize) + +lemma add_neg_numeral_simps: + "numeral m + neg_numeral n = sub m n" + "neg_numeral m + numeral n = sub n m" + "neg_numeral m + neg_numeral n = neg_numeral (m + n)" + unfolding sub_def diff_minus neg_numeral_def numeral_add numeral.simps + by (simp_all add: is_num_normalize) + +lemma add_neg_numeral_special: + "1 + neg_numeral m = sub One m" + "neg_numeral m + 1 = sub One m" + unfolding sub_def diff_minus neg_numeral_def numeral_add numeral.simps + by (simp_all add: is_num_normalize) + +lemma diff_numeral_simps: + "numeral m - numeral n = sub m n" + "numeral m - neg_numeral n = numeral (m + n)" + "neg_numeral m - numeral n = neg_numeral (m + n)" + "neg_numeral m - neg_numeral n = sub n m" + unfolding neg_numeral_def sub_def diff_minus numeral_add numeral.simps + by (simp_all add: is_num_normalize) + +lemma diff_numeral_special: + "1 - numeral n = sub One n" + "1 - neg_numeral n = numeral (One + n)" + "numeral m - 1 = sub m One" + "neg_numeral m - 1 = neg_numeral (m + One)" + unfolding neg_numeral_def sub_def diff_minus numeral_add numeral.simps + by (simp_all add: is_num_normalize) + +lemma minus_one: "- 1 = -1" + unfolding neg_numeral_def numeral.simps .. + +lemma minus_numeral: "- numeral n = neg_numeral n" + unfolding neg_numeral_def .. + +lemma minus_neg_numeral: "- neg_numeral n = numeral n" + unfolding neg_numeral_def by simp + +lemmas minus_numeral_simps [simp] = + minus_one minus_numeral minus_neg_numeral + +end + +subsubsection {* + Structures with multiplication: class @{text semiring_numeral} +*} + +class semiring_numeral = semiring + monoid_mult +begin + +subclass numeral .. + +lemma numeral_mult: "numeral (m * n) = numeral m * numeral n" + apply (induct n rule: num_induct) + apply (simp add: numeral_One) + apply (simp add: mult_inc numeral_inc numeral_add numeral_inc right_distrib) + done + +lemma numeral_times_numeral: "numeral m * numeral n = numeral (m * n)" + by (rule numeral_mult [symmetric]) + +end + +subsubsection {* + Structures with a zero: class @{text semiring_1} +*} + +context semiring_1 +begin + +subclass semiring_numeral .. + +lemma of_nat_numeral [simp]: "of_nat (numeral n) = numeral n" + by (induct n, + simp_all only: numeral.simps numeral_class.numeral.simps of_nat_add of_nat_1) + +end + +lemma nat_of_num_numeral: "nat_of_num = numeral" +proof + fix n + have "numeral n = nat_of_num n" + by (induct n) (simp_all add: numeral.simps) + then show "nat_of_num n = numeral n" by simp +qed + +subsubsection {* + Equality: class @{text semiring_char_0} +*} + +context semiring_char_0 +begin + +lemma numeral_eq_iff: "numeral m = numeral n \ m = n" + unfolding of_nat_numeral [symmetric] nat_of_num_numeral [symmetric] + of_nat_eq_iff num_eq_iff .. + +lemma numeral_eq_one_iff: "numeral n = 1 \ n = One" + by (rule numeral_eq_iff [of n One, unfolded numeral_One]) + +lemma one_eq_numeral_iff: "1 = numeral n \ One = n" + by (rule numeral_eq_iff [of One n, unfolded numeral_One]) + +lemma numeral_neq_zero: "numeral n \ 0" + unfolding of_nat_numeral [symmetric] nat_of_num_numeral [symmetric] + by (simp add: nat_of_num_pos) + +lemma zero_neq_numeral: "0 \ numeral n" + unfolding eq_commute [of 0] by (rule numeral_neq_zero) + +lemmas eq_numeral_simps [simp] = + numeral_eq_iff + numeral_eq_one_iff + one_eq_numeral_iff + numeral_neq_zero + zero_neq_numeral + +end + +subsubsection {* + Comparisons: class @{text linordered_semidom} +*} + +text {* Could be perhaps more general than here. *} + +context linordered_semidom +begin + +lemma numeral_le_iff: "numeral m \ numeral n \ m \ n" +proof - + have "of_nat (numeral m) \ of_nat (numeral n) \ m \ n" + unfolding less_eq_num_def nat_of_num_numeral of_nat_le_iff .. + then show ?thesis by simp +qed + +lemma one_le_numeral: "1 \ numeral n" +using numeral_le_iff [of One n] by (simp add: numeral_One) + +lemma numeral_le_one_iff: "numeral n \ 1 \ n \ One" +using numeral_le_iff [of n One] by (simp add: numeral_One) + +lemma numeral_less_iff: "numeral m < numeral n \ m < n" +proof - + have "of_nat (numeral m) < of_nat (numeral n) \ m < n" + unfolding less_num_def nat_of_num_numeral of_nat_less_iff .. + then show ?thesis by simp +qed + +lemma not_numeral_less_one: "\ numeral n < 1" + using numeral_less_iff [of n One] by (simp add: numeral_One) + +lemma one_less_numeral_iff: "1 < numeral n \ One < n" + using numeral_less_iff [of One n] by (simp add: numeral_One) + +lemma zero_le_numeral: "0 \ numeral n" + by (induct n) (simp_all add: numeral.simps) + +lemma zero_less_numeral: "0 < numeral n" + by (induct n) (simp_all add: numeral.simps add_pos_pos) + +lemma not_numeral_le_zero: "\ numeral n \ 0" + by (simp add: not_le zero_less_numeral) + +lemma not_numeral_less_zero: "\ numeral n < 0" + by (simp add: not_less zero_le_numeral) + +lemmas le_numeral_extra = + zero_le_one not_one_le_zero + order_refl [of 0] order_refl [of 1] + +lemmas less_numeral_extra = + zero_less_one not_one_less_zero + less_irrefl [of 0] less_irrefl [of 1] + +lemmas le_numeral_simps [simp] = + numeral_le_iff + one_le_numeral + numeral_le_one_iff + zero_le_numeral + not_numeral_le_zero + +lemmas less_numeral_simps [simp] = + numeral_less_iff + one_less_numeral_iff + not_numeral_less_one + zero_less_numeral + not_numeral_less_zero + +end + +subsubsection {* + Multiplication and negation: class @{text ring_1} +*} + +context ring_1 +begin + +subclass neg_numeral .. + +lemma mult_neg_numeral_simps: + "neg_numeral m * neg_numeral n = numeral (m * n)" + "neg_numeral m * numeral n = neg_numeral (m * n)" + "numeral m * neg_numeral n = neg_numeral (m * n)" + unfolding neg_numeral_def mult_minus_left mult_minus_right + by (simp_all only: minus_minus numeral_mult) + +lemma mult_minus1 [simp]: "-1 * z = - z" + unfolding neg_numeral_def numeral.simps mult_minus_left by simp + +lemma mult_minus1_right [simp]: "z * -1 = - z" + unfolding neg_numeral_def numeral.simps mult_minus_right by simp + +end + +subsubsection {* + Equality using @{text iszero} for rings with non-zero characteristic +*} + +context ring_1 +begin + +definition iszero :: "'a \ bool" + where "iszero z \ z = 0" + +lemma iszero_0 [simp]: "iszero 0" + by (simp add: iszero_def) + +lemma not_iszero_1 [simp]: "\ iszero 1" + by (simp add: iszero_def) + +lemma not_iszero_Numeral1: "\ iszero Numeral1" + by (simp add: numeral_One) + +lemma iszero_neg_numeral [simp]: + "iszero (neg_numeral w) \ iszero (numeral w)" + unfolding iszero_def neg_numeral_def + by (rule neg_equal_0_iff_equal) + +lemma eq_iff_iszero_diff: "x = y \ iszero (x - y)" + unfolding iszero_def by (rule eq_iff_diff_eq_0) + +text {* The @{text "eq_numeral_iff_iszero"} lemmas are not declared +@{text "[simp]"} by default, because for rings of characteristic zero, +better simp rules are possible. For a type like integers mod @{text +"n"}, type-instantiated versions of these rules should be added to the +simplifier, along with a type-specific rule for deciding propositions +of the form @{text "iszero (numeral w)"}. + +bh: Maybe it would not be so bad to just declare these as simp +rules anyway? I should test whether these rules take precedence over +the @{text "ring_char_0"} rules in the simplifier. +*} + +lemma eq_numeral_iff_iszero: + "numeral x = numeral y \ iszero (sub x y)" + "numeral x = neg_numeral y \ iszero (numeral (x + y))" + "neg_numeral x = numeral y \ iszero (numeral (x + y))" + "neg_numeral x = neg_numeral y \ iszero (sub y x)" + "numeral x = 1 \ iszero (sub x One)" + "1 = numeral y \ iszero (sub One y)" + "neg_numeral x = 1 \ iszero (numeral (x + One))" + "1 = neg_numeral y \ iszero (numeral (One + y))" + "numeral x = 0 \ iszero (numeral x)" + "0 = numeral y \ iszero (numeral y)" + "neg_numeral x = 0 \ iszero (numeral x)" + "0 = neg_numeral y \ iszero (numeral y)" + unfolding eq_iff_iszero_diff diff_numeral_simps diff_numeral_special + by simp_all + +end + +subsubsection {* + Equality and negation: class @{text ring_char_0} +*} + +class ring_char_0 = ring_1 + semiring_char_0 +begin + +lemma not_iszero_numeral [simp]: "\ iszero (numeral w)" + by (simp add: iszero_def) + +lemma neg_numeral_eq_iff: "neg_numeral m = neg_numeral n \ m = n" + by (simp only: neg_numeral_def neg_equal_iff_equal numeral_eq_iff) + +lemma numeral_neq_neg_numeral: "numeral m \ neg_numeral n" + unfolding neg_numeral_def eq_neg_iff_add_eq_0 + by (simp add: numeral_plus_numeral) + +lemma neg_numeral_neq_numeral: "neg_numeral m \ numeral n" + by (rule numeral_neq_neg_numeral [symmetric]) + +lemma zero_neq_neg_numeral: "0 \ neg_numeral n" + unfolding neg_numeral_def neg_0_equal_iff_equal by simp + +lemma neg_numeral_neq_zero: "neg_numeral n \ 0" + unfolding neg_numeral_def neg_equal_0_iff_equal by simp + +lemma one_neq_neg_numeral: "1 \ neg_numeral n" + using numeral_neq_neg_numeral [of One n] by (simp add: numeral_One) + +lemma neg_numeral_neq_one: "neg_numeral n \ 1" + using neg_numeral_neq_numeral [of n One] by (simp add: numeral_One) + +lemmas eq_neg_numeral_simps [simp] = + neg_numeral_eq_iff + numeral_neq_neg_numeral neg_numeral_neq_numeral + one_neq_neg_numeral neg_numeral_neq_one + zero_neq_neg_numeral neg_numeral_neq_zero + +end + +subsubsection {* + Structures with negation and order: class @{text linordered_idom} +*} + +context linordered_idom +begin + +subclass ring_char_0 .. + +lemma neg_numeral_le_iff: "neg_numeral m \ neg_numeral n \ n \ m" + by (simp only: neg_numeral_def neg_le_iff_le numeral_le_iff) + +lemma neg_numeral_less_iff: "neg_numeral m < neg_numeral n \ n < m" + by (simp only: neg_numeral_def neg_less_iff_less numeral_less_iff) + +lemma neg_numeral_less_zero: "neg_numeral n < 0" + by (simp only: neg_numeral_def neg_less_0_iff_less zero_less_numeral) + +lemma neg_numeral_le_zero: "neg_numeral n \ 0" + by (simp only: neg_numeral_def neg_le_0_iff_le zero_le_numeral) + +lemma not_zero_less_neg_numeral: "\ 0 < neg_numeral n" + by (simp only: not_less neg_numeral_le_zero) + +lemma not_zero_le_neg_numeral: "\ 0 \ neg_numeral n" + by (simp only: not_le neg_numeral_less_zero) + +lemma neg_numeral_less_numeral: "neg_numeral m < numeral n" + using neg_numeral_less_zero zero_less_numeral by (rule less_trans) + +lemma neg_numeral_le_numeral: "neg_numeral m \ numeral n" + by (simp only: less_imp_le neg_numeral_less_numeral) + +lemma not_numeral_less_neg_numeral: "\ numeral m < neg_numeral n" + by (simp only: not_less neg_numeral_le_numeral) + +lemma not_numeral_le_neg_numeral: "\ numeral m \ neg_numeral n" + by (simp only: not_le neg_numeral_less_numeral) + +lemma neg_numeral_less_one: "neg_numeral m < 1" + by (rule neg_numeral_less_numeral [of m One, unfolded numeral_One]) + +lemma neg_numeral_le_one: "neg_numeral m \ 1" + by (rule neg_numeral_le_numeral [of m One, unfolded numeral_One]) + +lemma not_one_less_neg_numeral: "\ 1 < neg_numeral m" + by (simp only: not_less neg_numeral_le_one) + +lemma not_one_le_neg_numeral: "\ 1 \ neg_numeral m" + by (simp only: not_le neg_numeral_less_one) + +lemma sub_non_negative: + "sub n m \ 0 \ n \ m" + by (simp only: sub_def le_diff_eq) simp + +lemma sub_positive: + "sub n m > 0 \ n > m" + by (simp only: sub_def less_diff_eq) simp + +lemma sub_non_positive: + "sub n m \ 0 \ n \ m" + by (simp only: sub_def diff_le_eq) simp + +lemma sub_negative: + "sub n m < 0 \ n < m" + by (simp only: sub_def diff_less_eq) simp + +lemmas le_neg_numeral_simps [simp] = + neg_numeral_le_iff + neg_numeral_le_numeral not_numeral_le_neg_numeral + neg_numeral_le_zero not_zero_le_neg_numeral + neg_numeral_le_one not_one_le_neg_numeral + +lemmas less_neg_numeral_simps [simp] = + neg_numeral_less_iff + neg_numeral_less_numeral not_numeral_less_neg_numeral + neg_numeral_less_zero not_zero_less_neg_numeral + neg_numeral_less_one not_one_less_neg_numeral + +lemma abs_numeral [simp]: "abs (numeral n) = numeral n" + by simp + +lemma abs_neg_numeral [simp]: "abs (neg_numeral n) = numeral n" + by (simp only: neg_numeral_def abs_minus_cancel abs_numeral) + +end + +subsubsection {* + Natural numbers +*} + +lemma Suc_numeral [simp]: "Suc (numeral n) = numeral (n + One)" + unfolding numeral_plus_one [symmetric] by simp + +lemma nat_number: + "1 = Suc 0" + "numeral One = Suc 0" + "numeral (Bit0 n) = Suc (numeral (BitM n))" + "numeral (Bit1 n) = Suc (numeral (Bit0 n))" + by (simp_all add: numeral.simps BitM_plus_one) + +subsubsection {* + Structures with exponentiation +*} + +context semiring_numeral +begin + +lemma numeral_sqr: "numeral (sqr n) = numeral n * numeral n" + by (simp add: sqr_conv_mult numeral_mult) + +lemma numeral_pow: "numeral (pow m n) = numeral m ^ numeral n" + by (induct n, simp_all add: numeral_class.numeral.simps + power_add numeral_sqr numeral_mult) + +lemma power_numeral [simp]: "numeral m ^ numeral n = numeral (pow m n)" + by (rule numeral_pow [symmetric]) + +end + +context semiring_1 +begin + +lemma power_zero_numeral [simp]: "(0::'a) ^ numeral n = 0" + by (induct n, simp_all add: numeral_class.numeral.simps power_add) + +end + +context ring_1 +begin + +lemma power_minus_Bit0: "(- x) ^ numeral (Bit0 n) = x ^ numeral (Bit0 n)" + by (induct n, simp_all add: numeral_class.numeral.simps power_add) + +lemma power_minus_Bit1: "(- x) ^ numeral (Bit1 n) = - (x ^ numeral (Bit1 n))" + by (simp only: nat_number(4) power_Suc power_minus_Bit0 mult_minus_left) + +lemma power_neg_numeral_Bit0 [simp]: + "neg_numeral m ^ numeral (Bit0 n) = numeral (pow m (Bit0 n))" + by (simp only: neg_numeral_def power_minus_Bit0 power_numeral) + +lemma power_neg_numeral_Bit1 [simp]: + "neg_numeral m ^ numeral (Bit1 n) = neg_numeral (pow m (Bit1 n))" + by (simp only: neg_numeral_def power_minus_Bit1 power_numeral pow.simps) + +end + +subsection {* Numeral equations as default simplification rules *} + +declare (in numeral) numeral_One [simp] +declare (in numeral) numeral_plus_numeral [simp] +declare (in numeral) add_numeral_special [simp] +declare (in neg_numeral) add_neg_numeral_simps [simp] +declare (in neg_numeral) add_neg_numeral_special [simp] +declare (in neg_numeral) diff_numeral_simps [simp] +declare (in neg_numeral) diff_numeral_special [simp] +declare (in semiring_numeral) numeral_times_numeral [simp] +declare (in ring_1) mult_neg_numeral_simps [simp] + +subsection {* Setting up simprocs *} + +lemma numeral_reorient: + "(numeral w = x) = (x = numeral w)" + by auto + +lemma mult_numeral_1: "Numeral1 * a = (a::'a::semiring_numeral)" + by simp + +lemma mult_numeral_1_right: "a * Numeral1 = (a::'a::semiring_numeral)" + by simp + +lemma divide_numeral_1: "a / Numeral1 = (a::'a::field)" + by simp + +lemma inverse_numeral_1: + "inverse Numeral1 = (Numeral1::'a::division_ring)" + by simp + +text{*Theorem lists for the cancellation simprocs. The use of a numary +numeral for 1 reduces the number of special cases.*} + +lemmas mult_1s = + mult_numeral_1 mult_numeral_1_right + mult_minus1 mult_minus1_right + + +subsubsection {* Simplification of arithmetic operations on integer constants. *} + +lemmas arith_special = (* already declared simp above *) + add_numeral_special add_neg_numeral_special + diff_numeral_special minus_one + +(* rules already in simpset *) +lemmas arith_extra_simps = + numeral_plus_numeral add_neg_numeral_simps add_0_left add_0_right + minus_numeral minus_neg_numeral minus_zero minus_one + diff_numeral_simps diff_0 diff_0_right + numeral_times_numeral mult_neg_numeral_simps + mult_zero_left mult_zero_right + abs_numeral abs_neg_numeral + +text {* + For making a minimal simpset, one must include these default simprules. + Also include @{text simp_thms}. +*} + +lemmas arith_simps = + add_num_simps mult_num_simps sub_num_simps + BitM.simps dbl_simps dbl_inc_simps dbl_dec_simps + abs_zero abs_one arith_extra_simps + +text {* Simplification of relational operations *} + +lemmas eq_numeral_extra = + zero_neq_one one_neq_zero + +lemmas rel_simps = + le_num_simps less_num_simps eq_num_simps + le_numeral_simps le_neg_numeral_simps le_numeral_extra + less_numeral_simps less_neg_numeral_simps less_numeral_extra + eq_numeral_simps eq_neg_numeral_simps eq_numeral_extra + + +subsubsection {* Simplification of arithmetic when nested to the right. *} + +lemma add_numeral_left [simp]: + "numeral v + (numeral w + z) = (numeral(v + w) + z)" + by (simp_all add: add_assoc [symmetric]) + +lemma add_neg_numeral_left [simp]: + "numeral v + (neg_numeral w + y) = (sub v w + y)" + "neg_numeral v + (numeral w + y) = (sub w v + y)" + "neg_numeral v + (neg_numeral w + y) = (neg_numeral(v + w) + y)" + by (simp_all add: add_assoc [symmetric]) + +lemma mult_numeral_left [simp]: + "numeral v * (numeral w * z) = (numeral(v * w) * z :: 'a::semiring_numeral)" + "neg_numeral v * (numeral w * y) = (neg_numeral(v * w) * y :: 'b::ring_1)" + "numeral v * (neg_numeral w * y) = (neg_numeral(v * w) * y :: 'b::ring_1)" + "neg_numeral v * (neg_numeral w * y) = (numeral(v * w) * y :: 'b::ring_1)" + by (simp_all add: mult_assoc [symmetric]) + +hide_const (open) One Bit0 Bit1 BitM inc pow sqr sub dbl dbl_inc dbl_dec + +subsection {* code module namespace *} + +code_modulename SML + Numeral Arith + +code_modulename OCaml + Numeral Arith + +code_modulename Haskell + Numeral Arith + +end diff -r 29e92b644d6c -r f067afe98049 src/HOL/Number_Theory/Primes.thy --- a/src/HOL/Number_Theory/Primes.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Number_Theory/Primes.thy Mon Mar 26 10:56:56 2012 +0200 @@ -206,7 +206,7 @@ "prime (p::nat) \ p > 1 \ (\n \ set [2.. n dvd p)" by (auto simp add: prime_nat_code) -lemmas prime_nat_simp_number_of [simp] = prime_nat_simp [of "number_of m"] for m +lemmas prime_nat_simp_numeral [simp] = prime_nat_simp [of "numeral m"] for m lemma prime_int_code [code]: "prime (p::int) \ p > 1 \ (\n \ {1<.. p > 1 \ (\n \ set [2..p - 1]. ~ n dvd p)" by (auto simp add: prime_int_code) -lemmas prime_int_simp_number_of [simp] = prime_int_simp [of "number_of m"] for m +lemmas prime_int_simp_numeral [simp] = prime_int_simp [of "numeral m"] for m lemma two_is_prime_nat [simp]: "prime (2::nat)" by simp diff -r 29e92b644d6c -r f067afe98049 src/HOL/Numeral_Simprocs.thy --- a/src/HOL/Numeral_Simprocs.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Numeral_Simprocs.thy Mon Mar 26 10:56:56 2012 +0200 @@ -14,8 +14,8 @@ ("Tools/nat_numeral_simprocs.ML") begin -declare split_div [of _ _ "number_of k", arith_split] for k -declare split_mod [of _ _ "number_of k", arith_split] for k +declare split_div [of _ _ "numeral k", arith_split] for k +declare split_mod [of _ _ "numeral k", arith_split] for k text {* For @{text combine_numerals} *} @@ -98,72 +98,74 @@ ("(a::'a::comm_semiring_1_cancel) * b") = {* fn phi => Numeral_Simprocs.assoc_fold *} +(* TODO: see whether the type class can be generalized further *) simproc_setup int_combine_numerals - ("(i::'a::number_ring) + j" | "(i::'a::number_ring) - j") = + ("(i::'a::comm_ring_1) + j" | "(i::'a::comm_ring_1) - j") = {* fn phi => Numeral_Simprocs.combine_numerals *} simproc_setup field_combine_numerals - ("(i::'a::{field_inverse_zero,ring_char_0,number_ring}) + j" - |"(i::'a::{field_inverse_zero,ring_char_0,number_ring}) - j") = + ("(i::'a::{field_inverse_zero,ring_char_0}) + j" + |"(i::'a::{field_inverse_zero,ring_char_0}) - j") = {* fn phi => Numeral_Simprocs.field_combine_numerals *} simproc_setup inteq_cancel_numerals - ("(l::'a::number_ring) + m = n" - |"(l::'a::number_ring) = m + n" - |"(l::'a::number_ring) - m = n" - |"(l::'a::number_ring) = m - n" - |"(l::'a::number_ring) * m = n" - |"(l::'a::number_ring) = m * n" - |"- (l::'a::number_ring) = m" - |"(l::'a::number_ring) = - m") = + ("(l::'a::comm_ring_1) + m = n" + |"(l::'a::comm_ring_1) = m + n" + |"(l::'a::comm_ring_1) - m = n" + |"(l::'a::comm_ring_1) = m - n" + |"(l::'a::comm_ring_1) * m = n" + |"(l::'a::comm_ring_1) = m * n" + |"- (l::'a::comm_ring_1) = m" + |"(l::'a::comm_ring_1) = - m") = {* fn phi => Numeral_Simprocs.eq_cancel_numerals *} simproc_setup intless_cancel_numerals - ("(l::'a::{linordered_idom,number_ring}) + m < n" - |"(l::'a::{linordered_idom,number_ring}) < m + n" - |"(l::'a::{linordered_idom,number_ring}) - m < n" - |"(l::'a::{linordered_idom,number_ring}) < m - n" - |"(l::'a::{linordered_idom,number_ring}) * m < n" - |"(l::'a::{linordered_idom,number_ring}) < m * n" - |"- (l::'a::{linordered_idom,number_ring}) < m" - |"(l::'a::{linordered_idom,number_ring}) < - m") = + ("(l::'a::linordered_idom) + m < n" + |"(l::'a::linordered_idom) < m + n" + |"(l::'a::linordered_idom) - m < n" + |"(l::'a::linordered_idom) < m - n" + |"(l::'a::linordered_idom) * m < n" + |"(l::'a::linordered_idom) < m * n" + |"- (l::'a::linordered_idom) < m" + |"(l::'a::linordered_idom) < - m") = {* fn phi => Numeral_Simprocs.less_cancel_numerals *} simproc_setup intle_cancel_numerals - ("(l::'a::{linordered_idom,number_ring}) + m \ n" - |"(l::'a::{linordered_idom,number_ring}) \ m + n" - |"(l::'a::{linordered_idom,number_ring}) - m \ n" - |"(l::'a::{linordered_idom,number_ring}) \ m - n" - |"(l::'a::{linordered_idom,number_ring}) * m \ n" - |"(l::'a::{linordered_idom,number_ring}) \ m * n" - |"- (l::'a::{linordered_idom,number_ring}) \ m" - |"(l::'a::{linordered_idom,number_ring}) \ - m") = + ("(l::'a::linordered_idom) + m \ n" + |"(l::'a::linordered_idom) \ m + n" + |"(l::'a::linordered_idom) - m \ n" + |"(l::'a::linordered_idom) \ m - n" + |"(l::'a::linordered_idom) * m \ n" + |"(l::'a::linordered_idom) \ m * n" + |"- (l::'a::linordered_idom) \ m" + |"(l::'a::linordered_idom) \ - m") = {* fn phi => Numeral_Simprocs.le_cancel_numerals *} simproc_setup ring_eq_cancel_numeral_factor - ("(l::'a::{idom,ring_char_0,number_ring}) * m = n" - |"(l::'a::{idom,ring_char_0,number_ring}) = m * n") = + ("(l::'a::{idom,ring_char_0}) * m = n" + |"(l::'a::{idom,ring_char_0}) = m * n") = {* fn phi => Numeral_Simprocs.eq_cancel_numeral_factor *} simproc_setup ring_less_cancel_numeral_factor - ("(l::'a::{linordered_idom,number_ring}) * m < n" - |"(l::'a::{linordered_idom,number_ring}) < m * n") = + ("(l::'a::linordered_idom) * m < n" + |"(l::'a::linordered_idom) < m * n") = {* fn phi => Numeral_Simprocs.less_cancel_numeral_factor *} simproc_setup ring_le_cancel_numeral_factor - ("(l::'a::{linordered_idom,number_ring}) * m <= n" - |"(l::'a::{linordered_idom,number_ring}) <= m * n") = + ("(l::'a::linordered_idom) * m <= n" + |"(l::'a::linordered_idom) <= m * n") = {* fn phi => Numeral_Simprocs.le_cancel_numeral_factor *} +(* TODO: remove comm_ring_1 constraint if possible *) simproc_setup int_div_cancel_numeral_factors - ("((l::'a::{semiring_div,ring_char_0,number_ring}) * m) div n" - |"(l::'a::{semiring_div,ring_char_0,number_ring}) div (m * n)") = + ("((l::'a::{semiring_div,comm_ring_1,ring_char_0}) * m) div n" + |"(l::'a::{semiring_div,comm_ring_1,ring_char_0}) div (m * n)") = {* fn phi => Numeral_Simprocs.div_cancel_numeral_factor *} simproc_setup divide_cancel_numeral_factor - ("((l::'a::{field_inverse_zero,ring_char_0,number_ring}) * m) / n" - |"(l::'a::{field_inverse_zero,ring_char_0,number_ring}) / (m * n)" - |"((number_of v)::'a::{field_inverse_zero,ring_char_0,number_ring}) / (number_of w)") = + ("((l::'a::{field_inverse_zero,ring_char_0}) * m) / n" + |"(l::'a::{field_inverse_zero,ring_char_0}) / (m * n)" + |"((numeral v)::'a::{field_inverse_zero,ring_char_0}) / (numeral w)") = {* fn phi => Numeral_Simprocs.divide_cancel_numeral_factor *} simproc_setup ring_eq_cancel_factor @@ -270,19 +272,25 @@ ("((l::nat) * m) dvd n" | "(l::nat) dvd (m * n)") = {* fn phi => Nat_Numeral_Simprocs.dvd_cancel_factor *} +(* FIXME: duplicate rule warnings for: + ring_distribs + numeral_plus_numeral numeral_times_numeral + numeral_eq_iff numeral_less_iff numeral_le_iff + numeral_neq_zero zero_neq_numeral zero_less_numeral + if_True if_False *) declaration {* - K (Lin_Arith.add_simps (@{thms neg_simps} @ [@{thm Suc_nat_number_of}, @{thm int_nat_number_of}]) - #> Lin_Arith.add_simps (@{thms ring_distribs} @ [@{thm Let_number_of}, @{thm Let_0}, @{thm Let_1}, + K (Lin_Arith.add_simps ([@{thm Suc_numeral}, @{thm int_numeral}]) + #> Lin_Arith.add_simps (@{thms ring_distribs} @ [@{thm Let_numeral}, @{thm Let_neg_numeral}, @{thm Let_0}, @{thm Let_1}, @{thm nat_0}, @{thm nat_1}, - @{thm add_nat_number_of}, @{thm diff_nat_number_of}, @{thm mult_nat_number_of}, - @{thm eq_nat_number_of}, @{thm less_nat_number_of}, @{thm le_number_of_eq_not_less}, - @{thm le_Suc_number_of}, @{thm le_number_of_Suc}, - @{thm less_Suc_number_of}, @{thm less_number_of_Suc}, - @{thm Suc_eq_number_of}, @{thm eq_number_of_Suc}, + @{thm numeral_plus_numeral}, @{thm diff_nat_numeral}, @{thm numeral_times_numeral}, + @{thm numeral_eq_iff}, @{thm numeral_less_iff}, @{thm numeral_le_iff}, + @{thm le_Suc_numeral}, @{thm le_numeral_Suc}, + @{thm less_Suc_numeral}, @{thm less_numeral_Suc}, + @{thm Suc_eq_numeral}, @{thm eq_numeral_Suc}, @{thm mult_Suc}, @{thm mult_Suc_right}, @{thm add_Suc}, @{thm add_Suc_right}, - @{thm eq_number_of_0}, @{thm eq_0_number_of}, @{thm less_0_number_of}, - @{thm of_int_number_of_eq}, @{thm of_nat_number_of_eq}, @{thm nat_number_of}, + @{thm numeral_neq_zero}, @{thm zero_neq_numeral}, @{thm zero_less_numeral}, + @{thm of_int_numeral}, @{thm of_nat_numeral}, @{thm nat_numeral}, @{thm if_True}, @{thm if_False}]) #> Lin_Arith.add_simprocs [@{simproc semiring_assoc_fold}, diff -r 29e92b644d6c -r f067afe98049 src/HOL/Parity.thy --- a/src/HOL/Parity.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Parity.thy Mon Mar 26 10:56:56 2012 +0200 @@ -45,9 +45,11 @@ lemma odd_1_nat [simp]: "odd (1::nat)" by presburger -declare even_def[of "number_of v", simp] for v +(* TODO: proper simp rules for Num.Bit0, Num.Bit1 *) +declare even_def[of "numeral v", simp] for v +declare even_def[of "neg_numeral v", simp] for v -declare even_nat_def[of "number_of v", simp] for v +declare even_nat_def[of "numeral v", simp] for v subsection {* Even and odd are mutually exclusive *} @@ -197,18 +199,18 @@ using minus_one_even_odd_power by blast lemma neg_one_even_odd_power: - "(even x --> (-1::'a::{number_ring})^x = 1) & + "(even x --> (-1::'a::{comm_ring_1})^x = 1) & (odd x --> (-1::'a)^x = -1)" apply (induct x) apply (simp, simp) done lemma neg_one_even_power [simp]: - "even x ==> (-1::'a::{number_ring})^x = 1" + "even x ==> (-1::'a::{comm_ring_1})^x = 1" using neg_one_even_odd_power by blast lemma neg_one_odd_power [simp]: - "odd x ==> (-1::'a::{number_ring})^x = -1" + "odd x ==> (-1::'a::{comm_ring_1})^x = -1" using neg_one_even_odd_power by blast lemma neg_power_if: @@ -347,27 +349,28 @@ text {* Simplify, when the exponent is a numeral *} -lemmas power_0_left_number_of = power_0_left [of "number_of w"] for w -declare power_0_left_number_of [simp] +lemma power_0_left_numeral [simp]: + "0 ^ numeral w = (0::'a::{power,semiring_0})" +by (simp add: power_0_left) -lemmas zero_le_power_eq_number_of [simp] = - zero_le_power_eq [of _ "number_of w"] for w +lemmas zero_le_power_eq_numeral [simp] = + zero_le_power_eq [of _ "numeral w"] for w -lemmas zero_less_power_eq_number_of [simp] = - zero_less_power_eq [of _ "number_of w"] for w +lemmas zero_less_power_eq_numeral [simp] = + zero_less_power_eq [of _ "numeral w"] for w -lemmas power_le_zero_eq_number_of [simp] = - power_le_zero_eq [of _ "number_of w"] for w +lemmas power_le_zero_eq_numeral [simp] = + power_le_zero_eq [of _ "numeral w"] for w -lemmas power_less_zero_eq_number_of [simp] = - power_less_zero_eq [of _ "number_of w"] for w +lemmas power_less_zero_eq_numeral [simp] = + power_less_zero_eq [of _ "numeral w"] for w -lemmas zero_less_power_nat_eq_number_of [simp] = - zero_less_power_nat_eq [of _ "number_of w"] for w +lemmas zero_less_power_nat_eq_numeral [simp] = + zero_less_power_nat_eq [of _ "numeral w"] for w -lemmas power_eq_0_iff_number_of [simp] = power_eq_0_iff [of _ "number_of w"] for w +lemmas power_eq_0_iff_numeral [simp] = power_eq_0_iff [of _ "numeral w"] for w -lemmas power_even_abs_number_of [simp] = power_even_abs [of "number_of w" _] for w +lemmas power_even_abs_numeral [simp] = power_even_abs [of "numeral w" _] for w subsection {* An Equivalence for @{term [source] "0 \ a^n"} *} diff -r 29e92b644d6c -r f067afe98049 src/HOL/Plain.thy --- a/src/HOL/Plain.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Plain.thy Mon Mar 26 10:56:56 2012 +0200 @@ -1,7 +1,7 @@ header {* Plain HOL *} theory Plain -imports Datatype FunDef Extraction Metis +imports Datatype FunDef Extraction Metis Num begin text {* diff -r 29e92b644d6c -r f067afe98049 src/HOL/Predicate_Compile_Examples/Predicate_Compile_Tests.thy --- a/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Tests.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Tests.thy Mon Mar 26 10:56:56 2012 +0200 @@ -334,7 +334,7 @@ code_pred [dseq] one_or_two . code_pred [random_dseq] one_or_two . thm one_or_two.dseq_equation -values [expected "{Suc 0::nat, 2::nat}"] "{x. one_or_two x}" +values [expected "{1::nat, 2::nat}"] "{x. one_or_two x}" values [random_dseq 0,0,10] 3 "{x. one_or_two x}" inductive one_or_two' :: "nat => bool" @@ -442,7 +442,7 @@ values "{ys. append [0, Suc 0, 2] ys [0, Suc 0, 2, 17, 0, 5]}" values [expected "{}" dseq 0] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}" -values [expected "{(([]::nat list), [Suc 0, 2, 3, 4, (5::nat)])}" dseq 1] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}" +values [expected "{(([]::nat list), [1, 2, 3, 4, (5::nat)])}" dseq 1] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}" values [dseq 4] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}" values [dseq 6] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}" values [random_dseq 1, 1, 4] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}" @@ -1241,8 +1241,8 @@ values [expected "{2::nat}"] "{x. plus_nat_test x 7 9}" values [expected "{}"] "{x. plus_nat_test x 9 7}" values [expected "{(0::nat,0::nat)}"] "{(x, y). plus_nat_test x y 0}" -values [expected "{(0, Suc 0), (Suc 0, 0)}"] "{(x, y). plus_nat_test x y 1}" -values [expected "{(0, 5), (4, Suc 0), (3, 2), (2, 3), (Suc 0, 4), (5, 0)}"] +values [expected "{(0::nat, 1::nat), (1, 0)}"] "{(x, y). plus_nat_test x y 1}" +values [expected "{(0::nat, 5::nat), (4, 1), (3, 2), (2, 3), (1, 4), (5, 0)}"] "{(x, y). plus_nat_test x y 5}" inductive minus_nat_test :: "nat => nat => nat => bool" @@ -1259,7 +1259,7 @@ values [expected "{5::nat}"] "{z. minus_nat_test 7 2 z}" values [expected "{16::nat}"] "{x. minus_nat_test x 7 9}" values [expected "{16::nat}"] "{x. minus_nat_test x 9 7}" -values [expected "{0, Suc 0, 2, 3}"] "{x. minus_nat_test x 3 0}" +values [expected "{0::nat, 1, 2, 3}"] "{x. minus_nat_test x 3 0}" values [expected "{0::nat}"] "{x. minus_nat_test x 0 0}" subsection {* Examples on int *} diff -r 29e92b644d6c -r f067afe98049 src/HOL/Presburger.thy --- a/src/HOL/Presburger.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Presburger.thy Mon Mar 26 10:56:56 2012 +0200 @@ -374,18 +374,16 @@ ((y \ x \ P (int x - int y)) \ (x < y \ P 0))" by (cases "y \ x") (simp_all add: zdiff_int) -lemma number_of1: "(0::int) <= number_of n \ (0::int) <= number_of (Int.Bit0 n) \ (0::int) <= number_of (Int.Bit1 n)" -by simp - -lemma number_of2: "(0::int) <= Numeral0" by simp - text {* \medskip Specific instances of congruence rules, to prevent simplifier from looping. *} -theorem imp_le_cong: "(0 <= x \ P = P') \ (0 <= (x::int) \ P) = (0 <= x \ P')" by simp +theorem imp_le_cong: + "\x = x'; 0 \ x' \ P = P'\ \ (0 \ (x::int) \ P) = (0 \ x' \ P')" + by simp -theorem conj_le_cong: "(0 <= x \ P = P') \ (0 <= (x::int) \ P) = (0 <= x \ P')" +theorem conj_le_cong: + "\x = x'; 0 \ x' \ P = P'\ \ (0 \ (x::int) \ P) = (0 \ x' \ P')" by (simp cong: conj_cong) use "Tools/Qelim/cooper.ML" diff -r 29e92b644d6c -r f067afe98049 src/HOL/Quickcheck_Examples/Quickcheck_Narrowing_Examples.thy --- a/src/HOL/Quickcheck_Examples/Quickcheck_Narrowing_Examples.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Quickcheck_Examples/Quickcheck_Narrowing_Examples.thy Mon Mar 26 10:56:56 2012 +0200 @@ -79,15 +79,14 @@ quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample] oops -(* FIXME: integer has strange representation! *) lemma "rev xs = xs" quickcheck[tester = narrowing, finite_types = false, default_type = int, expect = counterexample] oops -(* + lemma "rev xs = xs" quickcheck[tester = narrowing, finite_types = true, expect = counterexample] oops -*) + subsection {* Simple examples with functions *} lemma "map f xs = map g xs" diff -r 29e92b644d6c -r f067afe98049 src/HOL/Quickcheck_Narrowing.thy --- a/src/HOL/Quickcheck_Narrowing.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Quickcheck_Narrowing.thy Mon Mar 26 10:56:56 2012 +0200 @@ -70,34 +70,15 @@ "HOL.equal k l \ HOL.equal (int_of k) (int_of l)" instance proof -qed (auto simp add: equal_code_int_def equal_int_def eq_int_refl) +qed (auto simp add: equal_code_int_def equal_int_def equal_int_refl) end -instantiation code_int :: number -begin - -definition - "number_of = of_int" - -instance .. - -end - -lemma int_of_number [simp]: - "int_of (number_of k) = number_of k" - by (simp add: number_of_code_int_def number_of_is_id) - - definition nat_of :: "code_int => nat" where "nat_of i = nat (int_of i)" - - -code_datatype "number_of \ int \ code_int" - -instantiation code_int :: "{minus, linordered_semidom, semiring_div, linorder}" +instantiation code_int :: "{minus, linordered_semidom, semiring_div, neg_numeral, linorder}" begin definition [simp, code del]: @@ -110,6 +91,9 @@ "n + m = of_int (int_of n + int_of m)" definition [simp, code del]: + "- n = of_int (- int_of n)" + +definition [simp, code del]: "n - m = of_int (int_of n - int_of m)" definition [simp, code del]: @@ -127,34 +111,43 @@ definition [simp, code del]: "n < m \ int_of n < int_of m" - instance proof qed (auto simp add: code_int left_distrib zmult_zless_mono2) end -lemma zero_code_int_code [code, code_unfold]: - "(0\code_int) = Numeral0" - by (simp add: number_of_code_int_def Pls_def) +lemma int_of_numeral [simp]: + "int_of (numeral k) = numeral k" + by (induct k) (simp_all only: numeral.simps plus_code_int_def + one_code_int_def of_int_inverse UNIV_I) + +definition Num :: "num \ code_int" + where [code_abbrev]: "Num = numeral" + +lemma [code_abbrev]: + "- numeral k = (neg_numeral k :: code_int)" + by (unfold neg_numeral_def) simp + +code_datatype "0::code_int" Num lemma one_code_int_code [code, code_unfold]: "(1\code_int) = Numeral1" - by (simp add: number_of_code_int_def Pls_def Bit1_def) + by (simp only: numeral.simps) -definition div_mod_code_int :: "code_int \ code_int \ code_int \ code_int" where - [code del]: "div_mod_code_int n m = (n div m, n mod m)" +definition div_mod :: "code_int \ code_int \ code_int \ code_int" where + [code del]: "div_mod n m = (n div m, n mod m)" lemma [code]: - "div_mod_code_int n m = (if m = 0 then (0, n) else (n div m, n mod m))" - unfolding div_mod_code_int_def by auto + "div_mod n m = (if m = 0 then (0, n) else (n div m, n mod m))" + unfolding div_mod_def by auto lemma [code]: - "n div m = fst (div_mod_code_int n m)" - unfolding div_mod_code_int_def by simp + "n div m = fst (div_mod n m)" + unfolding div_mod_def by simp lemma [code]: - "n mod m = snd (div_mod_code_int n m)" - unfolding div_mod_code_int_def by simp + "n mod m = snd (div_mod n m)" + unfolding div_mod_def by simp lemma int_of_code [code]: "int_of k = (if k = 0 then 0 @@ -172,9 +165,12 @@ code_instance code_numeral :: equal (Haskell_Quickcheck -) -setup {* fold (Numeral.add_code @{const_name number_code_int_inst.number_of_code_int} +setup {* fold (Numeral.add_code @{const_name Num} false Code_Printer.literal_numeral) ["Haskell_Quickcheck"] *} +code_type code_int + (Haskell_Quickcheck "Int") + code_const "0 \ code_int" (Haskell_Quickcheck "0") @@ -182,24 +178,23 @@ (Haskell_Quickcheck "1") code_const "minus \ code_int \ code_int \ code_int" - (Haskell_Quickcheck "(_/ -/ _)") + (Haskell_Quickcheck infixl 6 "-") -code_const div_mod_code_int +code_const div_mod (Haskell_Quickcheck "divMod") code_const "HOL.equal \ code_int \ code_int \ bool" (Haskell_Quickcheck infix 4 "==") -code_const "op \ \ code_int \ code_int \ bool" +code_const "less_eq \ code_int \ code_int \ bool" (Haskell_Quickcheck infix 4 "<=") -code_const "op < \ code_int \ code_int \ bool" +code_const "less \ code_int \ code_int \ bool" (Haskell_Quickcheck infix 4 "<") -code_type code_int - (Haskell_Quickcheck "Int") +code_abort of_int -code_abort of_int +hide_const (open) Num div_mod subsubsection {* Narrowing's deep representation of types and terms *} diff -r 29e92b644d6c -r f067afe98049 src/HOL/Quotient.thy --- a/src/HOL/Quotient.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Quotient.thy Mon Mar 26 10:56:56 2012 +0200 @@ -9,7 +9,8 @@ keywords "print_quotmaps" "print_quotients" "print_quotconsts" :: diag and "quotient_type" :: thy_goal and "/" and - "quotient_definition" :: thy_decl + "setup_lifting" :: thy_decl and + "quotient_definition" :: thy_goal uses ("Tools/Quotient/quotient_info.ML") ("Tools/Quotient/quotient_type.ML") @@ -79,6 +80,10 @@ shows "((op =) ===> (op =)) = (op =)" by (auto simp add: fun_eq_iff elim: fun_relE) +lemma fun_rel_eq_rel: + shows "((op =) ===> R) = (\f g. \x. R (f x) (g x))" + by (simp add: fun_rel_def) + subsection {* set map (vimage) and set relation *} definition "set_rel R xs ys \ \x y. R x y \ x \ xs \ y \ ys" @@ -133,6 +138,18 @@ unfolding Quotient_def by blast +lemma Quotient_refl1: + assumes a: "Quotient R Abs Rep" + shows "R r s \ R r r" + using a unfolding Quotient_def + by fast + +lemma Quotient_refl2: + assumes a: "Quotient R Abs Rep" + shows "R r s \ R s s" + using a unfolding Quotient_def + by fast + lemma Quotient_rel_rep: assumes a: "Quotient R Abs Rep" shows "R (Rep a) (Rep b) \ a = b" @@ -259,6 +276,15 @@ shows "R2 (f x) (g y)" using a by (auto elim: fun_relE) +lemma apply_rsp'': + assumes "Quotient R Abs Rep" + and "(R ===> S) f f" + shows "S (f (Rep x)) (f (Rep x))" +proof - + from assms(1) have "R (Rep x) (Rep x)" by (rule Quotient_rep_reflp) + then show ?thesis using assms(2) by (auto intro: apply_rsp') +qed + subsection {* lemmas for regularisation of ball and bex *} lemma ball_reg_eqv: @@ -675,6 +701,153 @@ end +subsection {* Quotient composition *} + +lemma OOO_quotient: + fixes R1 :: "'a \ 'a \ bool" + fixes Abs1 :: "'a \ 'b" and Rep1 :: "'b \ 'a" + fixes Abs2 :: "'b \ 'c" and Rep2 :: "'c \ 'b" + fixes R2' :: "'a \ 'a \ bool" + fixes R2 :: "'b \ 'b \ bool" + assumes R1: "Quotient R1 Abs1 Rep1" + assumes R2: "Quotient R2 Abs2 Rep2" + assumes Abs1: "\x y. R2' x y \ R1 x x \ R1 y y \ R2 (Abs1 x) (Abs1 y)" + assumes Rep1: "\x y. R2 x y \ R2' (Rep1 x) (Rep1 y)" + shows "Quotient (R1 OO R2' OO R1) (Abs2 \ Abs1) (Rep1 \ Rep2)" +apply (rule QuotientI) + apply (simp add: o_def Quotient_abs_rep [OF R2] Quotient_abs_rep [OF R1]) + apply simp + apply (rule_tac b="Rep1 (Rep2 a)" in pred_compI) + apply (rule Quotient_rep_reflp [OF R1]) + apply (rule_tac b="Rep1 (Rep2 a)" in pred_compI [rotated]) + apply (rule Quotient_rep_reflp [OF R1]) + apply (rule Rep1) + apply (rule Quotient_rep_reflp [OF R2]) + apply safe + apply (rename_tac x y) + apply (drule Abs1) + apply (erule Quotient_refl2 [OF R1]) + apply (erule Quotient_refl1 [OF R1]) + apply (drule Quotient_refl1 [OF R2], drule Rep1) + apply (subgoal_tac "R1 r (Rep1 (Abs1 x))") + apply (rule_tac b="Rep1 (Abs1 x)" in pred_compI, assumption) + apply (erule pred_compI) + apply (erule Quotient_symp [OF R1, THEN sympD]) + apply (rule Quotient_rel[symmetric, OF R1, THEN iffD2]) + apply (rule conjI, erule Quotient_refl1 [OF R1]) + apply (rule conjI, rule Quotient_rep_reflp [OF R1]) + apply (subst Quotient_abs_rep [OF R1]) + apply (erule Quotient_rel_abs [OF R1]) + apply (rename_tac x y) + apply (drule Abs1) + apply (erule Quotient_refl2 [OF R1]) + apply (erule Quotient_refl1 [OF R1]) + apply (drule Quotient_refl2 [OF R2], drule Rep1) + apply (subgoal_tac "R1 s (Rep1 (Abs1 y))") + apply (rule_tac b="Rep1 (Abs1 y)" in pred_compI, assumption) + apply (erule pred_compI) + apply (erule Quotient_symp [OF R1, THEN sympD]) + apply (rule Quotient_rel[symmetric, OF R1, THEN iffD2]) + apply (rule conjI, erule Quotient_refl2 [OF R1]) + apply (rule conjI, rule Quotient_rep_reflp [OF R1]) + apply (subst Quotient_abs_rep [OF R1]) + apply (erule Quotient_rel_abs [OF R1, THEN sym]) + apply simp + apply (rule Quotient_rel_abs [OF R2]) + apply (rule Quotient_rel_abs [OF R1, THEN ssubst], assumption) + apply (rule Quotient_rel_abs [OF R1, THEN subst], assumption) + apply (erule Abs1) + apply (erule Quotient_refl2 [OF R1]) + apply (erule Quotient_refl1 [OF R1]) + apply (rename_tac a b c d) + apply simp + apply (rule_tac b="Rep1 (Abs1 r)" in pred_compI) + apply (rule Quotient_rel[symmetric, OF R1, THEN iffD2]) + apply (rule conjI, erule Quotient_refl1 [OF R1]) + apply (simp add: Quotient_abs_rep [OF R1] Quotient_rep_reflp [OF R1]) + apply (rule_tac b="Rep1 (Abs1 s)" in pred_compI [rotated]) + apply (rule Quotient_rel[symmetric, OF R1, THEN iffD2]) + apply (simp add: Quotient_abs_rep [OF R1] Quotient_rep_reflp [OF R1]) + apply (erule Quotient_refl2 [OF R1]) + apply (rule Rep1) + apply (drule Abs1) + apply (erule Quotient_refl2 [OF R1]) + apply (erule Quotient_refl1 [OF R1]) + apply (drule Abs1) + apply (erule Quotient_refl2 [OF R1]) + apply (erule Quotient_refl1 [OF R1]) + apply (drule Quotient_rel_abs [OF R1]) + apply (drule Quotient_rel_abs [OF R1]) + apply (drule Quotient_rel_abs [OF R1]) + apply (drule Quotient_rel_abs [OF R1]) + apply simp + apply (rule Quotient_rel[symmetric, OF R2, THEN iffD2]) + apply simp +done + +lemma OOO_eq_quotient: + fixes R1 :: "'a \ 'a \ bool" + fixes Abs1 :: "'a \ 'b" and Rep1 :: "'b \ 'a" + fixes Abs2 :: "'b \ 'c" and Rep2 :: "'c \ 'b" + assumes R1: "Quotient R1 Abs1 Rep1" + assumes R2: "Quotient op= Abs2 Rep2" + shows "Quotient (R1 OOO op=) (Abs2 \ Abs1) (Rep1 \ Rep2)" +using assms +by (rule OOO_quotient) auto + +subsection {* Invariant *} + +definition invariant :: "('a \ bool) \ 'a \ 'a \ bool" + where "invariant R = (\x y. R x \ x = y)" + +lemma invariant_to_eq: + assumes "invariant P x y" + shows "x = y" +using assms by (simp add: invariant_def) + +lemma fun_rel_eq_invariant: + shows "((invariant R) ===> S) = (\f g. \x. R x \ S (f x) (g x))" +by (auto simp add: invariant_def fun_rel_def) + +lemma invariant_same_args: + shows "invariant P x x \ P x" +using assms by (auto simp add: invariant_def) + +lemma copy_type_to_Quotient: + assumes "type_definition Rep Abs UNIV" + shows "Quotient (op =) Abs Rep" +proof - + interpret type_definition Rep Abs UNIV by fact + from Abs_inject Rep_inverse show ?thesis by (auto intro!: QuotientI) +qed + +lemma copy_type_to_equivp: + fixes Abs :: "'a \ 'b" + and Rep :: "'b \ 'a" + assumes "type_definition Rep Abs (UNIV::'a set)" + shows "equivp (op=::'a\'a\bool)" +by (rule identity_equivp) + +lemma invariant_type_to_Quotient: + assumes "type_definition Rep Abs {x. P x}" + shows "Quotient (invariant P) Abs Rep" +proof - + interpret type_definition Rep Abs "{x. P x}" by fact + from Rep Abs_inject Rep_inverse show ?thesis by (auto intro!: QuotientI simp: invariant_def) +qed + +lemma invariant_type_to_part_equivp: + assumes "type_definition Rep Abs {x. P x}" + shows "part_equivp (invariant P)" +proof (intro part_equivpI) + interpret type_definition Rep Abs "{x. P x}" by fact + show "\x. invariant P x x" using Rep by (auto simp: invariant_def) +next + show "symp (invariant P)" by (auto intro: sympI simp: invariant_def) +next + show "transp (invariant P)" by (auto intro: transpI simp: invariant_def) +qed + subsection {* ML setup *} text {* Auxiliary data for the quotient package *} @@ -682,8 +855,7 @@ use "Tools/Quotient/quotient_info.ML" setup Quotient_Info.setup -declare [[map "fun" = fun_rel]] -declare [[map set = set_rel]] +declare [[map "fun" = (fun_rel, fun_quotient)]] lemmas [quot_thm] = fun_quotient lemmas [quot_respect] = quot_rel_rsp if_rsp o_rsp let_rsp id_rsp @@ -788,4 +960,6 @@ map_fun (infixr "--->" 55) and fun_rel (infixr "===>" 55) +hide_const (open) invariant + end diff -r 29e92b644d6c -r f067afe98049 src/HOL/Quotient_Examples/DList.thy --- a/src/HOL/Quotient_Examples/DList.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Quotient_Examples/DList.thy Mon Mar 26 10:56:56 2012 +0200 @@ -88,45 +88,32 @@ definition [simp]: "card_remdups = length \ remdups" definition [simp]: "foldr_remdups f xs e = foldr f (remdups xs) e" -lemma [quot_respect]: - "(dlist_eq) Nil Nil" - "(dlist_eq ===> op =) List.member List.member" - "(op = ===> dlist_eq ===> dlist_eq) Cons Cons" - "(op = ===> dlist_eq ===> dlist_eq) removeAll removeAll" - "(dlist_eq ===> op =) card_remdups card_remdups" - "(dlist_eq ===> op =) remdups remdups" - "(op = ===> dlist_eq ===> op =) foldr_remdups foldr_remdups" - "(op = ===> dlist_eq ===> dlist_eq) map map" - "(op = ===> dlist_eq ===> dlist_eq) filter filter" - by (auto intro!: fun_relI simp add: remdups_filter) - (metis (full_types) set_remdups remdups_eq_map_eq remdups_eq_member_eq)+ - quotient_definition empty where "empty :: 'a dlist" - is "Nil" + is "Nil" done quotient_definition insert where "insert :: 'a \ 'a dlist \ 'a dlist" - is "Cons" + is "Cons" by (metis (mono_tags) List.insert_def dlist_eq_def remdups.simps(2) set_remdups) quotient_definition "member :: 'a dlist \ 'a \ bool" - is "List.member" + is "List.member" by (metis dlist_eq_def remdups_eq_member_eq) quotient_definition foldr where "foldr :: ('a \ 'b \ 'b) \ 'a dlist \ 'b \ 'b" - is "foldr_remdups" + is "foldr_remdups" by auto quotient_definition "remove :: 'a \ 'a dlist \ 'a dlist" - is "removeAll" + is "removeAll" by force quotient_definition card where "card :: 'a dlist \ nat" - is "card_remdups" + is "card_remdups" by fastforce quotient_definition map where "map :: ('a \ 'b) \ 'a dlist \ 'b dlist" - is "List.map :: ('a \ 'b) \ 'a list \ 'b list" + is "List.map :: ('a \ 'b) \ 'a list \ 'b list" by (metis dlist_eq_def remdups_eq_map_eq) quotient_definition filter where "filter :: ('a \ bool) \ 'a dlist \ 'a dlist" - is "List.filter" + is "List.filter" by (metis dlist_eq_def remdups_filter) quotient_definition "list_of_dlist :: 'a dlist \ 'a list" - is "remdups" + is "remdups" by simp text {* lifted theorems *} diff -r 29e92b644d6c -r f067afe98049 src/HOL/Quotient_Examples/FSet.thy --- a/src/HOL/Quotient_Examples/FSet.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Quotient_Examples/FSet.thy Mon Mar 26 10:56:56 2012 +0200 @@ -179,140 +179,6 @@ by (rule quotient_compose_list_g, rule Quotient_fset, rule list_eq_equivp) - -subsection {* Respectfulness lemmas for list operations *} - -lemma list_equiv_rsp [quot_respect]: - shows "(op \ ===> op \ ===> op =) op \ op \" - by (auto intro!: fun_relI) - -lemma append_rsp [quot_respect]: - shows "(op \ ===> op \ ===> op \) append append" - by (auto intro!: fun_relI) - -lemma sub_list_rsp [quot_respect]: - shows "(op \ ===> op \ ===> op =) sub_list sub_list" - by (auto intro!: fun_relI) - -lemma member_rsp [quot_respect]: - shows "(op \ ===> op =) List.member List.member" -proof - fix x y assume "x \ y" - then show "List.member x = List.member y" - unfolding fun_eq_iff by simp -qed - -lemma nil_rsp [quot_respect]: - shows "(op \) Nil Nil" - by simp - -lemma cons_rsp [quot_respect]: - shows "(op = ===> op \ ===> op \) Cons Cons" - by (auto intro!: fun_relI) - -lemma map_rsp [quot_respect]: - shows "(op = ===> op \ ===> op \) map map" - by (auto intro!: fun_relI) - -lemma set_rsp [quot_respect]: - "(op \ ===> op =) set set" - by (auto intro!: fun_relI) - -lemma inter_list_rsp [quot_respect]: - shows "(op \ ===> op \ ===> op \) inter_list inter_list" - by (auto intro!: fun_relI) - -lemma removeAll_rsp [quot_respect]: - shows "(op = ===> op \ ===> op \) removeAll removeAll" - by (auto intro!: fun_relI) - -lemma diff_list_rsp [quot_respect]: - shows "(op \ ===> op \ ===> op \) diff_list diff_list" - by (auto intro!: fun_relI) - -lemma card_list_rsp [quot_respect]: - shows "(op \ ===> op =) card_list card_list" - by (auto intro!: fun_relI) - -lemma filter_rsp [quot_respect]: - shows "(op = ===> op \ ===> op \) filter filter" - by (auto intro!: fun_relI) - -lemma remdups_removeAll: (*FIXME move*) - "remdups (removeAll x xs) = remove1 x (remdups xs)" - by (induct xs) auto - -lemma member_commute_fold_once: - assumes "rsp_fold f" - and "x \ set xs" - shows "fold_once f xs = fold_once f (removeAll x xs) \ f x" -proof - - from assms have "fold f (remdups xs) = fold f (remove1 x (remdups xs)) \ f x" - by (auto intro!: fold_remove1_split elim: rsp_foldE) - then show ?thesis using `rsp_fold f` by (simp add: fold_once_fold_remdups remdups_removeAll) -qed - -lemma fold_once_set_equiv: - assumes "xs \ ys" - shows "fold_once f xs = fold_once f ys" -proof (cases "rsp_fold f") - case False then show ?thesis by simp -next - case True - then have "\x y. x \ set (remdups xs) \ y \ set (remdups xs) \ f x \ f y = f y \ f x" - by (rule rsp_foldE) - moreover from assms have "multiset_of (remdups xs) = multiset_of (remdups ys)" - by (simp add: set_eq_iff_multiset_of_remdups_eq) - ultimately have "fold f (remdups xs) = fold f (remdups ys)" - by (rule fold_multiset_equiv) - with True show ?thesis by (simp add: fold_once_fold_remdups) -qed - -lemma fold_once_rsp [quot_respect]: - shows "(op = ===> op \ ===> op =) fold_once fold_once" - unfolding fun_rel_def by (auto intro: fold_once_set_equiv) - -lemma concat_rsp_pre: - assumes a: "list_all2 op \ x x'" - and b: "x' \ y'" - and c: "list_all2 op \ y' y" - and d: "\x\set x. xa \ set x" - shows "\x\set y. xa \ set x" -proof - - obtain xb where e: "xb \ set x" and f: "xa \ set xb" using d by auto - have "\y. y \ set x' \ xb \ y" by (rule list_all2_find_element[OF e a]) - then obtain ya where h: "ya \ set x'" and i: "xb \ ya" by auto - have "ya \ set y'" using b h by simp - then have "\yb. yb \ set y \ ya \ yb" using c by (rule list_all2_find_element) - then show ?thesis using f i by auto -qed - -lemma concat_rsp [quot_respect]: - shows "(list_all2 op \ OOO op \ ===> op \) concat concat" -proof (rule fun_relI, elim pred_compE) - fix a b ba bb - assume a: "list_all2 op \ a ba" - with list_symp [OF list_eq_symp] have a': "list_all2 op \ ba a" by (rule sympE) - assume b: "ba \ bb" - with list_eq_symp have b': "bb \ ba" by (rule sympE) - assume c: "list_all2 op \ bb b" - with list_symp [OF list_eq_symp] have c': "list_all2 op \ b bb" by (rule sympE) - have "\x. (\xa\set a. x \ set xa) = (\xa\set b. x \ set xa)" - proof - fix x - show "(\xa\set a. x \ set xa) = (\xa\set b. x \ set xa)" - proof - assume d: "\xa\set a. x \ set xa" - show "\xa\set b. x \ set xa" by (rule concat_rsp_pre[OF a b c d]) - next - assume e: "\xa\set b. x \ set xa" - show "\xa\set a. x \ set xa" by (rule concat_rsp_pre[OF c' b' a' e]) - qed - qed - then show "concat a \ concat b" by auto -qed - - section {* Quotient definitions for fsets *} @@ -323,7 +189,7 @@ quotient_definition "bot :: 'a fset" - is "Nil :: 'a list" + is "Nil :: 'a list" done abbreviation empty_fset ("{||}") @@ -332,7 +198,7 @@ quotient_definition "less_eq_fset :: ('a fset \ 'a fset \ bool)" - is "sub_list :: ('a list \ 'a list \ bool)" + is "sub_list :: ('a list \ 'a list \ bool)" by simp abbreviation subset_fset :: "'a fset \ 'a fset \ bool" (infix "|\|" 50) @@ -351,7 +217,7 @@ quotient_definition "sup :: 'a fset \ 'a fset \ 'a fset" - is "append :: 'a list \ 'a list \ 'a list" + is "append :: 'a list \ 'a list \ 'a list" by simp abbreviation union_fset (infixl "|\|" 65) @@ -360,7 +226,7 @@ quotient_definition "inf :: 'a fset \ 'a fset \ 'a fset" - is "inter_list :: 'a list \ 'a list \ 'a list" + is "inter_list :: 'a list \ 'a list \ 'a list" by simp abbreviation inter_fset (infixl "|\|" 65) @@ -369,7 +235,7 @@ quotient_definition "minus :: 'a fset \ 'a fset \ 'a fset" - is "diff_list :: 'a list \ 'a list \ 'a list" + is "diff_list :: 'a list \ 'a list \ 'a list" by fastforce instance proof @@ -413,7 +279,7 @@ quotient_definition "insert_fset :: 'a \ 'a fset \ 'a fset" - is "Cons" + is "Cons" by auto syntax "_insert_fset" :: "args => 'a fset" ("{|(_)|}") @@ -425,7 +291,7 @@ quotient_definition fset_member where - "fset_member :: 'a fset \ 'a \ bool" is "List.member" + "fset_member :: 'a fset \ 'a \ bool" is "List.member" by fastforce abbreviation in_fset :: "'a \ 'a fset \ bool" (infix "|\|" 50) @@ -442,31 +308,84 @@ quotient_definition "card_fset :: 'a fset \ nat" - is card_list + is card_list by simp quotient_definition "map_fset :: ('a \ 'b) \ 'a fset \ 'b fset" - is map + is map by simp quotient_definition "remove_fset :: 'a \ 'a fset \ 'a fset" - is removeAll + is removeAll by simp quotient_definition "fset :: 'a fset \ 'a set" - is "set" + is "set" by simp + +lemma fold_once_set_equiv: + assumes "xs \ ys" + shows "fold_once f xs = fold_once f ys" +proof (cases "rsp_fold f") + case False then show ?thesis by simp +next + case True + then have "\x y. x \ set (remdups xs) \ y \ set (remdups xs) \ f x \ f y = f y \ f x" + by (rule rsp_foldE) + moreover from assms have "multiset_of (remdups xs) = multiset_of (remdups ys)" + by (simp add: set_eq_iff_multiset_of_remdups_eq) + ultimately have "fold f (remdups xs) = fold f (remdups ys)" + by (rule fold_multiset_equiv) + with True show ?thesis by (simp add: fold_once_fold_remdups) +qed quotient_definition "fold_fset :: ('a \ 'b \ 'b) \ 'a fset \ 'b \ 'b" - is fold_once + is fold_once by (rule fold_once_set_equiv) + +lemma concat_rsp_pre: + assumes a: "list_all2 op \ x x'" + and b: "x' \ y'" + and c: "list_all2 op \ y' y" + and d: "\x\set x. xa \ set x" + shows "\x\set y. xa \ set x" +proof - + obtain xb where e: "xb \ set x" and f: "xa \ set xb" using d by auto + have "\y. y \ set x' \ xb \ y" by (rule list_all2_find_element[OF e a]) + then obtain ya where h: "ya \ set x'" and i: "xb \ ya" by auto + have "ya \ set y'" using b h by simp + then have "\yb. yb \ set y \ ya \ yb" using c by (rule list_all2_find_element) + then show ?thesis using f i by auto +qed quotient_definition "concat_fset :: ('a fset) fset \ 'a fset" - is concat + is concat +proof (elim pred_compE) +fix a b ba bb + assume a: "list_all2 op \ a ba" + with list_symp [OF list_eq_symp] have a': "list_all2 op \ ba a" by (rule sympE) + assume b: "ba \ bb" + with list_eq_symp have b': "bb \ ba" by (rule sympE) + assume c: "list_all2 op \ bb b" + with list_symp [OF list_eq_symp] have c': "list_all2 op \ b bb" by (rule sympE) + have "\x. (\xa\set a. x \ set xa) = (\xa\set b. x \ set xa)" + proof + fix x + show "(\xa\set a. x \ set xa) = (\xa\set b. x \ set xa)" + proof + assume d: "\xa\set a. x \ set xa" + show "\xa\set b. x \ set xa" by (rule concat_rsp_pre[OF a b c d]) + next + assume e: "\xa\set b. x \ set xa" + show "\xa\set a. x \ set xa" by (rule concat_rsp_pre[OF c' b' a' e]) + qed + qed + then show "concat a \ concat b" by auto +qed quotient_definition "filter_fset :: ('a \ bool) \ 'a fset \ 'a fset" - is filter + is filter by force subsection {* Compositional respectfulness and preservation lemmas *} @@ -538,7 +457,7 @@ lemma append_rsp2 [quot_respect]: "(list_all2 op \ OOO op \ ===> list_all2 op \ OOO op \ ===> list_all2 op \ OOO op \) append append" - by (intro compositional_rsp3 append_rsp) + by (intro compositional_rsp3) (auto intro!: fun_relI simp add: append_rsp2_pre) lemma map_rsp2 [quot_respect]: @@ -967,6 +886,20 @@ (if rsp_fold f then if a |\| A then fold_fset f A else fold_fset f A \ f a else id)" by descending (simp add: fold_once_fold_remdups) +lemma remdups_removeAll: + "remdups (removeAll x xs) = remove1 x (remdups xs)" + by (induct xs) auto + +lemma member_commute_fold_once: + assumes "rsp_fold f" + and "x \ set xs" + shows "fold_once f xs = fold_once f (removeAll x xs) \ f x" +proof - + from assms have "fold f (remdups xs) = fold f (remove1 x (remdups xs)) \ f x" + by (auto intro!: fold_remove1_split elim: rsp_foldE) + then show ?thesis using `rsp_fold f` by (simp add: fold_once_fold_remdups remdups_removeAll) +qed + lemma in_commute_fold_fset: "rsp_fold f \ h |\| b \ fold_fset f b = fold_fset f (remove_fset h b) \ f h" by descending (simp add: member_commute_fold_once) @@ -1170,7 +1103,7 @@ then have e': "List.member r a" using list_eq_def [simplified List.member_def [symmetric], of l r] b by auto have f: "card_list (removeAll a l) = m" using e d by (simp) - have g: "removeAll a l \ removeAll a r" using removeAll_rsp b by simp + have g: "removeAll a l \ removeAll a r" using remove_fset_rsp b by simp have "(removeAll a l) \2 (removeAll a r)" by (rule Suc.hyps[OF f g]) then have h: "(a # removeAll a l) \2 (a # removeAll a r)" by (rule list_eq2.intros(5)) have i: "l \2 (a # removeAll a l)" diff -r 29e92b644d6c -r f067afe98049 src/HOL/Quotient_Examples/Lift_Fun.thy --- a/src/HOL/Quotient_Examples/Lift_Fun.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Quotient_Examples/Lift_Fun.thy Mon Mar 26 10:56:56 2012 +0200 @@ -6,7 +6,7 @@ theory Lift_Fun -imports Main +imports Main "~~/src/HOL/Library/Quotient_Syntax" begin text {* This file is meant as a test case for features introduced in the changeset 2d8949268303. @@ -23,17 +23,17 @@ by (simp add: identity_equivp) quotient_definition "comp' :: ('b \ 'c) \ ('a \ 'b) \ 'a \ 'c" is - "comp :: ('b \ 'c) \ ('a \ 'b) \ 'a \ 'c" + "comp :: ('b \ 'c) \ ('a \ 'b) \ 'a \ 'c" done quotient_definition "fcomp' :: ('a \ 'b) \ ('b \ 'c) \ 'a \ 'c" is - fcomp + fcomp done quotient_definition "map_fun' :: ('c \ 'a) \ ('b \ 'd) \ ('a \ 'b) \ 'c \ 'd" - is "map_fun::('c \ 'a) \ ('b \ 'd) \ ('a \ 'b) \ 'c \ 'd" + is "map_fun::('c \ 'a) \ ('b \ 'd) \ ('a \ 'b) \ 'c \ 'd" done -quotient_definition "inj_on' :: ('a \ 'b) \ 'a set \ bool" is inj_on +quotient_definition "inj_on' :: ('a \ 'b) \ 'a set \ bool" is inj_on done -quotient_definition "bij_betw' :: ('a \ 'b) \ 'a set \ 'b set \ bool" is bij_betw +quotient_definition "bij_betw' :: ('a \ 'b) \ 'a set \ 'b set \ bool" is bij_betw done subsection {* Co/Contravariant type variables *} @@ -47,7 +47,7 @@ where "map_endofun' f g e = map_fun g f e" quotient_definition "map_endofun :: ('a \ 'b) \ ('b \ 'a) \ 'a endofun \ 'b endofun" is - map_endofun' + map_endofun' done text {* Registration of the map function for 'a endofun. *} @@ -63,7 +63,45 @@ by (auto simp: map_endofun_def map_endofun'_def map_fun_def fun_eq_iff) (simp add: a o_assoc) qed -quotient_definition "endofun_id_id :: ('a endofun) endofun" is "id :: ('a \ 'a) \ ('a \ 'a)" +text {* Relator for 'a endofun. *} + +definition + endofun_rel' :: "('a \ 'b \ bool) \ ('a \ 'a) \ ('b \ 'b) \ bool" +where + "endofun_rel' R = (\f g. (R ===> R) f g)" + +quotient_definition "endofun_rel :: ('a \ 'b \ bool) \ 'a endofun \ 'b endofun \ bool" is + endofun_rel' done + +lemma endofun_quotient: +assumes a: "Quotient R Abs Rep" +shows "Quotient (endofun_rel R) (map_endofun Abs Rep) (map_endofun Rep Abs)" +proof (intro QuotientI) + show "\a. map_endofun Abs Rep (map_endofun Rep Abs a) = a" + by (metis (hide_lams, no_types) a abs_o_rep id_apply map_endofun.comp map_endofun.id o_eq_dest_lhs) +next + show "\a. endofun_rel R (map_endofun Rep Abs a) (map_endofun Rep Abs a)" + using fun_quotient[OF a a, THEN Quotient_rep_reflp] + unfolding endofun_rel_def map_endofun_def map_fun_def o_def map_endofun'_def endofun_rel'_def id_def + by (metis (mono_tags) Quotient_endofun rep_abs_rsp) +next + show "\r s. endofun_rel R r s = + (endofun_rel R r r \ + endofun_rel R s s \ map_endofun Abs Rep r = map_endofun Abs Rep s)" + apply(auto simp add: endofun_rel_def endofun_rel'_def map_endofun_def map_endofun'_def) + using fun_quotient[OF a a,THEN Quotient_refl1] + apply metis + using fun_quotient[OF a a,THEN Quotient_refl2] + apply metis + using fun_quotient[OF a a, THEN Quotient_rel] + apply metis + using fun_quotient[OF a a, THEN Quotient_rel] + by (smt Quotient_endofun rep_abs_rsp) +qed + +declare [[map endofun = (endofun_rel, endofun_quotient)]] + +quotient_definition "endofun_id_id :: ('a endofun) endofun" is "id :: ('a \ 'a) \ ('a \ 'a)" done term endofun_id_id thm endofun_id_id_def @@ -73,7 +111,7 @@ text {* We have to map "'a endofun" to "('a endofun') endofun", i.e., mapping (lifting) over a type variable which is a covariant and contravariant type variable. *} -quotient_definition "endofun'_id_id :: ('a endofun') endofun'" is endofun_id_id +quotient_definition "endofun'_id_id :: ('a endofun') endofun'" is endofun_id_id done term endofun'_id_id thm endofun'_id_id_def diff -r 29e92b644d6c -r f067afe98049 src/HOL/Quotient_Examples/Lift_RBT.thy --- a/src/HOL/Quotient_Examples/Lift_RBT.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Quotient_Examples/Lift_RBT.thy Mon Mar 26 10:56:56 2012 +0200 @@ -15,21 +15,6 @@ then show ?thesis .. qed -local_setup {* fn lthy => -let - val quotients = {qtyp = @{typ "('a, 'b) rbt"}, rtyp = @{typ "('a, 'b) RBT_Impl.rbt"}, - equiv_rel = @{term "dummy"}, equiv_thm = @{thm refl}} - val qty_full_name = @{type_name "rbt"} - - fun qinfo phi = Quotient_Info.transform_quotients phi quotients - in lthy - |> Local_Theory.declaration {syntax = false, pervasive = true} - (fn phi => Quotient_Info.update_quotients qty_full_name (qinfo phi) - #> Quotient_Info.update_abs_rep qty_full_name (Quotient_Info.transform_abs_rep phi - {abs = @{term "RBT"}, rep = @{term "impl_of"}})) - end -*} - lemma rbt_eq_iff: "t1 = t2 \ impl_of t1 = impl_of t2" by (simp add: impl_of_inject) @@ -46,12 +31,12 @@ "RBT (impl_of t) = t" by (simp add: impl_of_inverse) - subsection {* Primitive operations *} -quotient_definition lookup where "lookup :: ('a\linorder, 'b) rbt \ 'a \ 'b" is "RBT_Impl.lookup" +setup_lifting type_definition_rbt -declare lookup_def[unfolded map_fun_def comp_def id_def, code] +quotient_definition lookup where "lookup :: ('a\linorder, 'b) rbt \ 'a \ 'b" is "RBT_Impl.lookup" +by simp (* FIXME: quotient_definition at the moment requires that types variables match exactly, i.e., sort constraints must be annotated to the constant being lifted. @@ -67,65 +52,38 @@ *) quotient_definition empty where "empty :: ('a\linorder, 'b) rbt" -is "(RBT_Impl.Empty :: ('a\linorder, 'b) RBT_Impl.rbt)" - -lemma impl_of_empty [code abstract]: - "impl_of empty = RBT_Impl.Empty" - by (simp add: empty_def RBT_inverse) +is "(RBT_Impl.Empty :: ('a\linorder, 'b) RBT_Impl.rbt)" by (simp add: empty_def) quotient_definition insert where "insert :: 'a\linorder \ 'b \ ('a, 'b) rbt \ ('a, 'b) rbt" -is "RBT_Impl.insert" - -lemma impl_of_insert [code abstract]: - "impl_of (insert k v t) = RBT_Impl.insert k v (impl_of t)" - by (simp add: insert_def RBT_inverse) +is "RBT_Impl.insert" by simp quotient_definition delete where "delete :: 'a\linorder \ ('a, 'b) rbt \ ('a, 'b) rbt" -is "RBT_Impl.delete" - -lemma impl_of_delete [code abstract]: - "impl_of (delete k t) = RBT_Impl.delete k (impl_of t)" - by (simp add: delete_def RBT_inverse) +is "RBT_Impl.delete" by simp (* FIXME: unnecessary type annotations *) quotient_definition "entries :: ('a\linorder, 'b) rbt \ ('a \ 'b) list" -is "RBT_Impl.entries :: ('a\linorder, 'b) RBT_Impl.rbt \ ('a \ 'b) list" - -lemma [code]: "entries t = RBT_Impl.entries (impl_of t)" -unfolding entries_def map_fun_def comp_def id_def .. +is "RBT_Impl.entries :: ('a\linorder, 'b) RBT_Impl.rbt \ ('a \ 'b) list" by simp (* FIXME: unnecessary type annotations *) quotient_definition "keys :: ('a\linorder, 'b) rbt \ 'a list" -is "RBT_Impl.keys :: ('a\linorder, 'b) RBT_Impl.rbt \ 'a list" +is "RBT_Impl.keys :: ('a\linorder, 'b) RBT_Impl.rbt \ 'a list" by simp quotient_definition "bulkload :: ('a\linorder \ 'b) list \ ('a, 'b) rbt" -is "RBT_Impl.bulkload" - -lemma impl_of_bulkload [code abstract]: - "impl_of (bulkload xs) = RBT_Impl.bulkload xs" - by (simp add: bulkload_def RBT_inverse) +is "RBT_Impl.bulkload" by simp quotient_definition "map_entry :: 'a \ ('b \ 'b) \ ('a\linorder, 'b) rbt \ ('a, 'b) rbt" -is "RBT_Impl.map_entry" - -lemma impl_of_map_entry [code abstract]: - "impl_of (map_entry k f t) = RBT_Impl.map_entry k f (impl_of t)" - by (simp add: map_entry_def RBT_inverse) +is "RBT_Impl.map_entry" by simp (* FIXME: unnecesary type annotations *) quotient_definition map where "map :: ('a \ 'b \ 'b) \ ('a\linorder, 'b) rbt \ ('a, 'b) rbt" is "RBT_Impl.map :: ('a \ 'b \ 'b) \ ('a\linorder, 'b) RBT_Impl.rbt \ ('a, 'b) RBT_Impl.rbt" - -lemma impl_of_map [code abstract]: - "impl_of (map f t) = RBT_Impl.map f (impl_of t)" - by (simp add: map_def RBT_inverse) +by simp (* FIXME: unnecessary type annotations *) -quotient_definition fold where "fold :: ('a \ 'b \ 'c \ 'c) \ ('a\linorder, 'b) rbt \ 'c \ 'c" is "RBT_Impl.fold :: ('a \ 'b \ 'c \ 'c) \ ('a\linorder, 'b) RBT_Impl.rbt \ 'c \ 'c" +quotient_definition fold where "fold :: ('a \ 'b \ 'c \ 'c) \ ('a\linorder, 'b) rbt \ 'c \ 'c" +is "RBT_Impl.fold :: ('a \ 'b \ 'c \ 'c) \ ('a\linorder, 'b) RBT_Impl.rbt \ 'c \ 'c" by simp -lemma [code]: "fold f t = RBT_Impl.fold f (impl_of t)" -unfolding fold_def map_fun_def comp_def id_def .. - +export_code lookup empty insert delete entries keys bulkload map_entry map fold in SML subsection {* Derived operations *} @@ -189,6 +147,10 @@ "fold f t = List.fold (prod_case f) (entries t)" by (simp add: fold_def fun_eq_iff RBT_Impl.fold_def entries_impl_of) +lemma impl_of_empty: + "impl_of empty = RBT_Impl.Empty" + by (simp add: empty_def RBT_inverse) + lemma is_empty_empty [simp]: "is_empty t \ t = empty" by (simp add: rbt_eq_iff is_empty_def impl_of_empty split: rbt.split) @@ -210,5 +172,4 @@ by (simp add: keys_def RBT_Impl.keys_def distinct_entries) - end \ No newline at end of file diff -r 29e92b644d6c -r f067afe98049 src/HOL/Quotient_Examples/Lift_Set.thy --- a/src/HOL/Quotient_Examples/Lift_Set.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Quotient_Examples/Lift_Set.thy Mon Mar 26 10:56:56 2012 +0200 @@ -14,30 +14,12 @@ morphisms member Set unfolding set_def by auto -text {* Here is some ML setup that should eventually be incorporated in the typedef command. *} - -local_setup {* fn lthy => - let - val quotients = - {qtyp = @{typ "'a set"}, rtyp = @{typ "'a => bool"}, - equiv_rel = @{term "dummy"}, equiv_thm = @{thm refl}} - val qty_full_name = @{type_name "set"} - - fun qinfo phi = Quotient_Info.transform_quotients phi quotients - in - lthy - |> Local_Theory.declaration {syntax = false, pervasive = true} - (fn phi => - Quotient_Info.update_quotients qty_full_name (qinfo phi) #> - Quotient_Info.update_abs_rep qty_full_name - (Quotient_Info.transform_abs_rep phi {abs = @{term "Set"}, rep = @{term "member"}})) - end -*} +setup_lifting type_definition_set[unfolded set_def] text {* Now, we can employ quotient_definition to lift definitions. *} quotient_definition empty where "empty :: 'a set" -is "bot :: 'a \ bool" +is "bot :: 'a \ bool" done term "Lift_Set.empty" thm Lift_Set.empty_def @@ -46,10 +28,12 @@ "insertp x P y \ y = x \ P y" quotient_definition insert where "insert :: 'a => 'a set => 'a set" -is insertp +is insertp done term "Lift_Set.insert" thm Lift_Set.insert_def +export_code empty insert in SML + end diff -r 29e92b644d6c -r f067afe98049 src/HOL/Quotient_Examples/Quotient_Cset.thy --- a/src/HOL/Quotient_Examples/Quotient_Cset.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Quotient_Examples/Quotient_Cset.thy Mon Mar 26 10:56:56 2012 +0200 @@ -21,75 +21,50 @@ subsection {* Operations *} -lemma [quot_respect]: - "(op = ===> set_eq ===> op =) (op \) (op \)" - "(op = ===> set_eq) Collect Collect" - "(set_eq ===> op =) Set.is_empty Set.is_empty" - "(op = ===> set_eq ===> set_eq) Set.insert Set.insert" - "(op = ===> set_eq ===> set_eq) Set.remove Set.remove" - "(op = ===> set_eq ===> set_eq) image image" - "(op = ===> set_eq ===> set_eq) Set.project Set.project" - "(set_eq ===> op =) Ball Ball" - "(set_eq ===> op =) Bex Bex" - "(set_eq ===> op =) Finite_Set.card Finite_Set.card" - "(set_eq ===> set_eq ===> op =) (op \) (op \)" - "(set_eq ===> set_eq ===> op =) (op \) (op \)" - "(set_eq ===> set_eq ===> set_eq) (op \) (op \)" - "(set_eq ===> set_eq ===> set_eq) (op \) (op \)" - "set_eq {} {}" - "set_eq UNIV UNIV" - "(set_eq ===> set_eq) uminus uminus" - "(set_eq ===> set_eq ===> set_eq) minus minus" - "(set_eq ===> op =) Inf Inf" - "(set_eq ===> op =) Sup Sup" - "(op = ===> set_eq) List.set List.set" - "(set_eq ===> (op = ===> set_eq) ===> set_eq) UNION UNION" -by (auto simp: fun_rel_eq) - quotient_definition "member :: 'a => 'a Quotient_Cset.set => bool" -is "op \" +is "op \" by simp quotient_definition "Set :: ('a => bool) => 'a Quotient_Cset.set" -is Collect +is Collect done quotient_definition is_empty where "is_empty :: 'a Quotient_Cset.set \ bool" -is Set.is_empty +is Set.is_empty by simp quotient_definition insert where "insert :: 'a \ 'a Quotient_Cset.set \ 'a Quotient_Cset.set" -is Set.insert +is Set.insert by simp quotient_definition remove where "remove :: 'a \ 'a Quotient_Cset.set \ 'a Quotient_Cset.set" -is Set.remove +is Set.remove by simp quotient_definition map where "map :: ('a \ 'b) \ 'a Quotient_Cset.set \ 'b Quotient_Cset.set" -is image +is image by simp quotient_definition filter where "filter :: ('a \ bool) \ 'a Quotient_Cset.set \ 'a Quotient_Cset.set" -is Set.project +is Set.project by simp quotient_definition "forall :: 'a Quotient_Cset.set \ ('a \ bool) \ bool" -is Ball +is Ball by simp quotient_definition "exists :: 'a Quotient_Cset.set \ ('a \ bool) \ bool" -is Bex +is Bex by simp quotient_definition card where "card :: 'a Quotient_Cset.set \ nat" -is Finite_Set.card +is Finite_Set.card by simp quotient_definition set where "set :: 'a list => 'a Quotient_Cset.set" -is List.set +is List.set done quotient_definition subset where "subset :: 'a Quotient_Cset.set \ 'a Quotient_Cset.set \ bool" -is "op \ :: 'a set \ 'a set \ bool" +is "op \ :: 'a set \ 'a set \ bool" by simp quotient_definition psubset where "psubset :: 'a Quotient_Cset.set \ 'a Quotient_Cset.set \ bool" -is "op \ :: 'a set \ 'a set \ bool" +is "op \ :: 'a set \ 'a set \ bool" by simp quotient_definition inter where "inter :: 'a Quotient_Cset.set \ 'a Quotient_Cset.set \ 'a Quotient_Cset.set" -is "op \ :: 'a set \ 'a set \ 'a set" +is "op \ :: 'a set \ 'a set \ 'a set" by simp quotient_definition union where "union :: 'a Quotient_Cset.set \ 'a Quotient_Cset.set \ 'a Quotient_Cset.set" -is "op \ :: 'a set \ 'a set \ 'a set" +is "op \ :: 'a set \ 'a set \ 'a set" by simp quotient_definition empty where "empty :: 'a Quotient_Cset.set" -is "{} :: 'a set" +is "{} :: 'a set" done quotient_definition UNIV where "UNIV :: 'a Quotient_Cset.set" -is "Set.UNIV :: 'a set" +is "Set.UNIV :: 'a set" done quotient_definition uminus where "uminus :: 'a Quotient_Cset.set \ 'a Quotient_Cset.set" -is "uminus_class.uminus :: 'a set \ 'a set" +is "uminus_class.uminus :: 'a set \ 'a set" by simp quotient_definition minus where "minus :: 'a Quotient_Cset.set \ 'a Quotient_Cset.set \ 'a Quotient_Cset.set" -is "(op -) :: 'a set \ 'a set \ 'a set" +is "(op -) :: 'a set \ 'a set \ 'a set" by simp quotient_definition Inf where "Inf :: ('a :: Inf) Quotient_Cset.set \ 'a" -is "Inf_class.Inf :: ('a :: Inf) set \ 'a" +is "Inf_class.Inf :: ('a :: Inf) set \ 'a" by simp quotient_definition Sup where "Sup :: ('a :: Sup) Quotient_Cset.set \ 'a" -is "Sup_class.Sup :: ('a :: Sup) set \ 'a" +is "Sup_class.Sup :: ('a :: Sup) set \ 'a" by simp quotient_definition UNION where "UNION :: 'a Quotient_Cset.set \ ('a \ 'b Quotient_Cset.set) \ 'b Quotient_Cset.set" -is "Complete_Lattices.UNION :: 'a set \ ('a \ 'b set) \ 'b set" +is "Complete_Lattices.UNION :: 'a set \ ('a \ 'b set) \ 'b set" by simp hide_const (open) is_empty insert remove map filter forall exists card set subset psubset inter union empty UNIV uminus minus Inf Sup UNION diff -r 29e92b644d6c -r f067afe98049 src/HOL/Quotient_Examples/Quotient_Int.thy --- a/src/HOL/Quotient_Examples/Quotient_Int.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Quotient_Examples/Quotient_Int.thy Mon Mar 26 10:56:56 2012 +0200 @@ -22,10 +22,10 @@ begin quotient_definition - "0 \ int" is "(0\nat, 0\nat)" + "0 \ int" is "(0\nat, 0\nat)" done quotient_definition - "1 \ int" is "(1\nat, 0\nat)" + "1 \ int" is "(1\nat, 0\nat)" done fun plus_int_raw :: "(nat \ nat) \ (nat \ nat) \ (nat \ nat)" @@ -33,7 +33,7 @@ "plus_int_raw (x, y) (u, v) = (x + u, y + v)" quotient_definition - "(op +) \ (int \ int \ int)" is "plus_int_raw" + "(op +) \ (int \ int \ int)" is "plus_int_raw" by auto fun uminus_int_raw :: "(nat \ nat) \ (nat \ nat)" @@ -41,7 +41,7 @@ "uminus_int_raw (x, y) = (y, x)" quotient_definition - "(uminus \ (int \ int))" is "uminus_int_raw" + "(uminus \ (int \ int))" is "uminus_int_raw" by auto definition minus_int_def: "z - w = z + (-w\int)" @@ -51,8 +51,38 @@ where "times_int_raw (x, y) (u, v) = (x*u + y*v, x*v + y*u)" +lemma times_int_raw_fst: + assumes a: "x \ z" + shows "times_int_raw x y \ times_int_raw z y" + using a + apply(cases x, cases y, cases z) + apply(auto simp add: times_int_raw.simps intrel.simps) + apply(rename_tac u v w x y z) + apply(subgoal_tac "u*w + z*w = y*w + v*w & u*x + z*x = y*x + v*x") + apply(simp add: mult_ac) + apply(simp add: add_mult_distrib [symmetric]) +done + +lemma times_int_raw_snd: + assumes a: "x \ z" + shows "times_int_raw y x \ times_int_raw y z" + using a + apply(cases x, cases y, cases z) + apply(auto simp add: times_int_raw.simps intrel.simps) + apply(rename_tac u v w x y z) + apply(subgoal_tac "u*w + z*w = y*w + v*w & u*x + z*x = y*x + v*x") + apply(simp add: mult_ac) + apply(simp add: add_mult_distrib [symmetric]) +done + quotient_definition "(op *) :: (int \ int \ int)" is "times_int_raw" + apply(rule equivp_transp[OF int_equivp]) + apply(rule times_int_raw_fst) + apply(assumption) + apply(rule times_int_raw_snd) + apply(assumption) +done fun le_int_raw :: "(nat \ nat) \ (nat \ nat) \ bool" @@ -60,7 +90,7 @@ "le_int_raw (x, y) (u, v) = (x+v \ u+y)" quotient_definition - le_int_def: "(op \) :: int \ int \ bool" is "le_int_raw" + le_int_def: "(op \) :: int \ int \ bool" is "le_int_raw" by auto definition less_int_def: "(z\int) < w = (z \ w \ z \ w)" @@ -75,47 +105,6 @@ end -lemma [quot_respect]: - shows "(op \ ===> op \ ===> op \) plus_int_raw plus_int_raw" - and "(op \ ===> op \) uminus_int_raw uminus_int_raw" - and "(op \ ===> op \ ===> op =) le_int_raw le_int_raw" - by (auto intro!: fun_relI) - -lemma times_int_raw_fst: - assumes a: "x \ z" - shows "times_int_raw x y \ times_int_raw z y" - using a - apply(cases x, cases y, cases z) - apply(auto simp add: times_int_raw.simps intrel.simps) - apply(rename_tac u v w x y z) - apply(subgoal_tac "u*w + z*w = y*w + v*w & u*x + z*x = y*x + v*x") - apply(simp add: mult_ac) - apply(simp add: add_mult_distrib [symmetric]) - done - -lemma times_int_raw_snd: - assumes a: "x \ z" - shows "times_int_raw y x \ times_int_raw y z" - using a - apply(cases x, cases y, cases z) - apply(auto simp add: times_int_raw.simps intrel.simps) - apply(rename_tac u v w x y z) - apply(subgoal_tac "u*w + z*w = y*w + v*w & u*x + z*x = y*x + v*x") - apply(simp add: mult_ac) - apply(simp add: add_mult_distrib [symmetric]) - done - -lemma [quot_respect]: - shows "(op \ ===> op \ ===> op \) times_int_raw times_int_raw" - apply(simp only: fun_rel_def) - apply(rule allI | rule impI)+ - apply(rule equivp_transp[OF int_equivp]) - apply(rule times_int_raw_fst) - apply(assumption) - apply(rule times_int_raw_snd) - apply(assumption) - done - text{* The integers form a @{text comm_ring_1}*} @@ -165,11 +154,7 @@ "int_of_nat_raw m = (m :: nat, 0 :: nat)" quotient_definition - "int_of_nat :: nat \ int" is "int_of_nat_raw" - -lemma[quot_respect]: - shows "(op = ===> op \) int_of_nat_raw int_of_nat_raw" - by (auto simp add: equivp_reflp [OF int_equivp]) + "int_of_nat :: nat \ int" is "int_of_nat_raw" done lemma int_of_nat: shows "of_nat m = int_of_nat m" @@ -304,11 +289,7 @@ quotient_definition "int_to_nat::int \ nat" is - "int_to_nat_raw" - -lemma [quot_respect]: - shows "(intrel ===> op =) int_to_nat_raw int_to_nat_raw" - by (auto iff: int_to_nat_raw_def) + "int_to_nat_raw" unfolding int_to_nat_raw_def by force lemma nat_le_eq_zle: fixes w z::"int" diff -r 29e92b644d6c -r f067afe98049 src/HOL/Quotient_Examples/Quotient_Message.thy --- a/src/HOL/Quotient_Examples/Quotient_Message.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Quotient_Examples/Quotient_Message.thy Mon Mar 26 10:56:56 2012 +0200 @@ -136,29 +136,25 @@ "Nonce :: nat \ msg" is "NONCE" +done quotient_definition "MPair :: msg \ msg \ msg" is "MPAIR" +by (rule MPAIR) quotient_definition "Crypt :: nat \ msg \ msg" is "CRYPT" +by (rule CRYPT) quotient_definition "Decrypt :: nat \ msg \ msg" is "DECRYPT" - -lemma [quot_respect]: - shows "(op = ===> op \ ===> op \) CRYPT CRYPT" -by (auto intro: CRYPT) - -lemma [quot_respect]: - shows "(op = ===> op \ ===> op \) DECRYPT DECRYPT" -by (auto intro: DECRYPT) +by (rule DECRYPT) text{*Establishing these two equations is the point of the whole exercise*} theorem CD_eq [simp]: @@ -175,25 +171,14 @@ "nonces:: msg \ nat set" is "freenonces" +by (rule msgrel_imp_eq_freenonces) text{*Now prove the four equations for @{term nonces}*} -lemma [quot_respect]: - shows "(op \ ===> op =) freenonces freenonces" - by (auto simp add: msgrel_imp_eq_freenonces) - -lemma [quot_respect]: - shows "(op = ===> op \) NONCE NONCE" - by (auto simp add: NONCE) - lemma nonces_Nonce [simp]: shows "nonces (Nonce N) = {N}" by (lifting freenonces.simps(1)) -lemma [quot_respect]: - shows " (op \ ===> op \ ===> op \) MPAIR MPAIR" - by (auto simp add: MPAIR) - lemma nonces_MPair [simp]: shows "nonces (MPair X Y) = nonces X \ nonces Y" by (lifting freenonces.simps(2)) @@ -212,10 +197,7 @@ "left:: msg \ msg" is "freeleft" - -lemma [quot_respect]: - shows "(op \ ===> op \) freeleft freeleft" - by (auto simp add: msgrel_imp_eqv_freeleft) +by (rule msgrel_imp_eqv_freeleft) lemma left_Nonce [simp]: shows "left (Nonce N) = Nonce N" @@ -239,13 +221,10 @@ "right:: msg \ msg" is "freeright" +by (rule msgrel_imp_eqv_freeright) text{*Now prove the four equations for @{term right}*} -lemma [quot_respect]: - shows "(op \ ===> op \) freeright freeright" - by (auto simp add: msgrel_imp_eqv_freeright) - lemma right_Nonce [simp]: shows "right (Nonce N) = Nonce N" by (lifting freeright.simps(1)) @@ -352,13 +331,10 @@ "discrim:: msg \ int" is "freediscrim" +by (rule msgrel_imp_eq_freediscrim) text{*Now prove the four equations for @{term discrim}*} -lemma [quot_respect]: - shows "(op \ ===> op =) freediscrim freediscrim" - by (auto simp add: msgrel_imp_eq_freediscrim) - lemma discrim_Nonce [simp]: shows "discrim (Nonce N) = 0" by (lifting freediscrim.simps(1)) diff -r 29e92b644d6c -r f067afe98049 src/HOL/Quotient_Examples/Quotient_Rat.thy --- a/src/HOL/Quotient_Examples/Quotient_Rat.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Quotient_Examples/Quotient_Rat.thy Mon Mar 26 10:56:56 2012 +0200 @@ -32,28 +32,29 @@ begin quotient_definition - "0 \ rat" is "(0\int, 1\int)" + "0 \ rat" is "(0\int, 1\int)" by simp quotient_definition - "1 \ rat" is "(1\int, 1\int)" + "1 \ rat" is "(1\int, 1\int)" by simp fun times_rat_raw where "times_rat_raw (a :: int, b :: int) (c, d) = (a * c, b * d)" quotient_definition - "(op *) :: (rat \ rat \ rat)" is times_rat_raw + "(op *) :: (rat \ rat \ rat)" is times_rat_raw by (auto simp add: mult_assoc mult_left_commute) fun plus_rat_raw where "plus_rat_raw (a :: int, b :: int) (c, d) = (a * d + c * b, b * d)" quotient_definition - "(op +) :: (rat \ rat \ rat)" is plus_rat_raw + "(op +) :: (rat \ rat \ rat)" is plus_rat_raw + by (auto simp add: mult_commute mult_left_commute int_distrib(2)) fun uminus_rat_raw where "uminus_rat_raw (a :: int, b :: int) = (-a, b)" quotient_definition - "(uminus \ (rat \ rat))" is "uminus_rat_raw" + "(uminus \ (rat \ rat))" is "uminus_rat_raw" by fastforce definition minus_rat_def: "a - b = a + (-b\rat)" @@ -63,6 +64,32 @@ quotient_definition "(op \) :: rat \ rat \ bool" is "le_rat_raw" +proof - + { + fix a b c d e f g h :: int + assume "a * f * (b * f) \ e * b * (b * f)" + then have le: "a * f * b * f \ e * b * b * f" by simp + assume nz: "b \ 0" "d \ 0" "f \ 0" "h \ 0" + then have b2: "b * b > 0" + by (metis linorder_neqE_linordered_idom mult_eq_0_iff not_square_less_zero) + have f2: "f * f > 0" using nz(3) + by (metis linorder_neqE_linordered_idom mult_eq_0_iff not_square_less_zero) + assume eq: "a * d = c * b" "e * h = g * f" + have "a * f * b * f * d * d \ e * b * b * f * d * d" using le nz(2) + by (metis linorder_le_cases mult_right_mono mult_right_mono_neg) + then have "c * f * f * d * (b * b) \ e * f * d * d * (b * b)" using eq + by (metis (no_types) mult_assoc mult_commute) + then have "c * f * f * d \ e * f * d * d" using b2 + by (metis leD linorder_le_less_linear mult_strict_right_mono) + then have "c * f * f * d * h * h \ e * f * d * d * h * h" using nz(4) + by (metis linorder_le_cases mult_right_mono mult_right_mono_neg) + then have "c * h * (d * h) * (f * f) \ g * d * (d * h) * (f * f)" using eq + by (metis (no_types) mult_assoc mult_commute) + then have "c * h * (d * h) \ g * d * (d * h)" using f2 + by (metis leD linorder_le_less_linear mult_strict_right_mono) + } + then show "\x y xa ya. x \ y \ xa \ ya \ le_rat_raw x xa = le_rat_raw y ya" by auto +qed definition less_rat_def: "(z\rat) < w = (z \ w \ z \ w)" @@ -83,14 +110,7 @@ where [simp]: "Fract_raw a b = (if b = 0 then (0, 1) else (a, b))" quotient_definition "Fract :: int \ int \ rat" is - Fract_raw - -lemma [quot_respect]: - "(op \ ===> op \ ===> op \) times_rat_raw times_rat_raw" - "(op \ ===> op \ ===> op \) plus_rat_raw plus_rat_raw" - "(op \ ===> op \) uminus_rat_raw uminus_rat_raw" - "(op = ===> op = ===> op \) Fract_raw Fract_raw" - by (auto intro!: fun_relI simp add: mult_assoc mult_commute mult_left_commute int_distrib(2)) + Fract_raw by simp lemmas [simp] = Respects_def @@ -139,32 +159,17 @@ apply auto done -instantiation rat :: number_ring -begin - -definition - rat_number_of_def: "number_of w = Fract w 1" - -instance apply default - unfolding rat_number_of_def of_int_rat .. - -end - instantiation rat :: field_inverse_zero begin fun rat_inverse_raw where "rat_inverse_raw (a :: int, b :: int) = (if a = 0 then (0, 1) else (b, a))" quotient_definition - "inverse :: rat \ rat" is rat_inverse_raw + "inverse :: rat \ rat" is rat_inverse_raw by (force simp add: mult_commute) definition divide_rat_def: "q / r = q * inverse (r::rat)" -lemma [quot_respect]: - "(op \ ===> op \) rat_inverse_raw rat_inverse_raw" - by (auto intro!: fun_relI simp add: mult_commute) - instance proof fix q :: rat assume "q \ 0" @@ -179,34 +184,6 @@ end -lemma [quot_respect]: "(op \ ===> op \ ===> op =) le_rat_raw le_rat_raw" -proof - - { - fix a b c d e f g h :: int - assume "a * f * (b * f) \ e * b * (b * f)" - then have le: "a * f * b * f \ e * b * b * f" by simp - assume nz: "b \ 0" "d \ 0" "f \ 0" "h \ 0" - then have b2: "b * b > 0" - by (metis linorder_neqE_linordered_idom mult_eq_0_iff not_square_less_zero) - have f2: "f * f > 0" using nz(3) - by (metis linorder_neqE_linordered_idom mult_eq_0_iff not_square_less_zero) - assume eq: "a * d = c * b" "e * h = g * f" - have "a * f * b * f * d * d \ e * b * b * f * d * d" using le nz(2) - by (metis linorder_le_cases mult_right_mono mult_right_mono_neg) - then have "c * f * f * d * (b * b) \ e * f * d * d * (b * b)" using eq - by (metis (no_types) mult_assoc mult_commute) - then have "c * f * f * d \ e * f * d * d" using b2 - by (metis leD linorder_le_less_linear mult_strict_right_mono) - then have "c * f * f * d * h * h \ e * f * d * d * h * h" using nz(4) - by (metis linorder_le_cases mult_right_mono mult_right_mono_neg) - then have "c * h * (d * h) * (f * f) \ g * d * (d * h) * (f * f)" using eq - by (metis (no_types) mult_assoc mult_commute) - then have "c * h * (d * h) \ g * d * (d * h)" using f2 - by (metis leD linorder_le_less_linear mult_strict_right_mono) - } - then show ?thesis by (auto intro!: fun_relI) -qed - instantiation rat :: linorder begin diff -r 29e92b644d6c -r f067afe98049 src/HOL/RComplete.thy --- a/src/HOL/RComplete.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/RComplete.thy Mon Mar 26 10:56:56 2012 +0200 @@ -129,26 +129,27 @@ subsection{*Floor and Ceiling Functions from the Reals to the Integers*} -lemma number_of_less_real_of_int_iff [simp]: - "((number_of n) < real (m::int)) = (number_of n < m)" +(* FIXME: theorems for negative numerals *) +lemma numeral_less_real_of_int_iff [simp]: + "((numeral n) < real (m::int)) = (numeral n < m)" apply auto apply (rule real_of_int_less_iff [THEN iffD1]) apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto) done -lemma number_of_less_real_of_int_iff2 [simp]: - "(real (m::int) < (number_of n)) = (m < number_of n)" +lemma numeral_less_real_of_int_iff2 [simp]: + "(real (m::int) < (numeral n)) = (m < numeral n)" apply auto apply (rule real_of_int_less_iff [THEN iffD1]) apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto) done -lemma number_of_le_real_of_int_iff [simp]: - "((number_of n) \ real (m::int)) = (number_of n \ m)" +lemma numeral_le_real_of_int_iff [simp]: + "((numeral n) \ real (m::int)) = (numeral n \ m)" by (simp add: linorder_not_less [symmetric]) -lemma number_of_le_real_of_int_iff2 [simp]: - "(real (m::int) \ (number_of n)) = (m \ number_of n)" +lemma numeral_le_real_of_int_iff2 [simp]: + "(real (m::int) \ (numeral n)) = (m \ numeral n)" by (simp add: linorder_not_less [symmetric]) lemma floor_real_of_nat [simp]: "floor (real (n::nat)) = int n" @@ -323,7 +324,7 @@ lemma zero_le_natfloor [simp]: "0 <= natfloor x" by (unfold natfloor_def, simp) -lemma natfloor_number_of_eq [simp]: "natfloor (number_of n) = number_of n" +lemma natfloor_numeral_eq [simp]: "natfloor (numeral n) = numeral n" by (unfold natfloor_def, simp) lemma natfloor_real_of_nat [simp]: "natfloor(real n) = n" @@ -365,9 +366,9 @@ apply (erule le_natfloor) done -lemma le_natfloor_eq_number_of [simp]: - "~ neg((number_of n)::int) ==> 0 <= x ==> - (number_of n <= natfloor x) = (number_of n <= x)" +lemma le_natfloor_eq_numeral [simp]: + "~ neg((numeral n)::int) ==> 0 <= x ==> + (numeral n <= natfloor x) = (numeral n <= x)" apply (subst le_natfloor_eq, assumption) apply simp done @@ -407,9 +408,9 @@ unfolding real_of_int_of_nat_eq [symmetric] floor_add by (simp add: nat_add_distrib) -lemma natfloor_add_number_of [simp]: - "~neg ((number_of n)::int) ==> 0 <= x ==> - natfloor (x + number_of n) = natfloor x + number_of n" +lemma natfloor_add_numeral [simp]: + "~neg ((numeral n)::int) ==> 0 <= x ==> + natfloor (x + numeral n) = natfloor x + numeral n" by (simp add: natfloor_add [symmetric]) lemma natfloor_add_one: "0 <= x ==> natfloor(x + 1) = natfloor x + 1" @@ -453,7 +454,7 @@ lemma zero_le_natceiling [simp]: "0 <= natceiling x" by (unfold natceiling_def, simp) -lemma natceiling_number_of_eq [simp]: "natceiling (number_of n) = number_of n" +lemma natceiling_numeral_eq [simp]: "natceiling (numeral n) = numeral n" by (unfold natceiling_def, simp) lemma natceiling_real_of_nat [simp]: "natceiling(real n) = n" @@ -476,9 +477,9 @@ unfolding natceiling_def real_of_nat_def by (simp add: nat_le_iff ceiling_le_iff) -lemma natceiling_le_eq_number_of [simp]: - "~ neg((number_of n)::int) ==> - (natceiling x <= number_of n) = (x <= number_of n)" +lemma natceiling_le_eq_numeral [simp]: + "~ neg((numeral n)::int) ==> + (natceiling x <= numeral n) = (x <= numeral n)" by (simp add: natceiling_le_eq) lemma natceiling_le_eq_one: "(natceiling x <= 1) = (x <= 1)" @@ -495,9 +496,9 @@ unfolding real_of_int_of_nat_eq [symmetric] ceiling_add by (simp add: nat_add_distrib) -lemma natceiling_add_number_of [simp]: - "~ neg ((number_of n)::int) ==> 0 <= x ==> - natceiling (x + number_of n) = natceiling x + number_of n" +lemma natceiling_add_numeral [simp]: + "~ neg ((numeral n)::int) ==> 0 <= x ==> + natceiling (x + numeral n) = natceiling x + numeral n" by (simp add: natceiling_add [symmetric]) lemma natceiling_add_one: "0 <= x ==> natceiling(x + 1) = natceiling x + 1" diff -r 29e92b644d6c -r f067afe98049 src/HOL/Rat.thy --- a/src/HOL/Rat.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Rat.thy Mon Mar 26 10:56:56 2012 +0200 @@ -230,35 +230,23 @@ lemma Fract_of_int_eq: "Fract k 1 = of_int k" by (rule of_int_rat [symmetric]) -instantiation rat :: number_ring -begin - -definition - rat_number_of_def: "number_of w = Fract w 1" - -instance proof -qed (simp add: rat_number_of_def of_int_rat) - -end - lemma rat_number_collapse: "Fract 0 k = 0" "Fract 1 1 = 1" - "Fract (number_of k) 1 = number_of k" + "Fract (numeral w) 1 = numeral w" + "Fract (neg_numeral w) 1 = neg_numeral w" "Fract k 0 = 0" - by (cases "k = 0") - (simp_all add: Zero_rat_def One_rat_def number_of_is_id number_of_eq of_int_rat eq_rat Fract_def) + using Fract_of_int_eq [of "numeral w"] + using Fract_of_int_eq [of "neg_numeral w"] + by (simp_all add: Zero_rat_def One_rat_def eq_rat) -lemma rat_number_expand [code_unfold]: +lemma rat_number_expand: "0 = Fract 0 1" "1 = Fract 1 1" - "number_of k = Fract (number_of k) 1" + "numeral k = Fract (numeral k) 1" + "neg_numeral k = Fract (neg_numeral k) 1" by (simp_all add: rat_number_collapse) -lemma iszero_rat [simp]: - "iszero (number_of k :: rat) \ iszero (number_of k :: int)" - by (simp add: iszero_def rat_number_expand number_of_is_id eq_rat) - lemma Rat_cases_nonzero [case_names Fract 0]: assumes Fract: "\a b. q = Fract a b \ b > 0 \ a \ 0 \ coprime a b \ C" assumes 0: "q = 0 \ C" @@ -386,7 +374,8 @@ lemma quotient_of_number [simp]: "quotient_of 0 = (0, 1)" "quotient_of 1 = (1, 1)" - "quotient_of (number_of k) = (number_of k, 1)" + "quotient_of (numeral k) = (numeral k, 1)" + "quotient_of (neg_numeral k) = (neg_numeral k, 1)" by (simp_all add: rat_number_expand quotient_of_Fract) lemma quotient_of_eq: "quotient_of (Fract a b) = (p, q) \ Fract p q = Fract a b" @@ -453,19 +442,12 @@ subsubsection {* Various *} -lemma Fract_add_one: "n \ 0 ==> Fract (m + n) n = Fract m n + 1" - by (simp add: rat_number_expand) - lemma Fract_of_int_quotient: "Fract k l = of_int k / of_int l" by (simp add: Fract_of_int_eq [symmetric]) -lemma Fract_number_of_quotient: - "Fract (number_of k) (number_of l) = number_of k / number_of l" - unfolding Fract_of_int_quotient number_of_is_id number_of_eq .. +lemma Fract_add_one: "n \ 0 ==> Fract (m + n) n = Fract m n + 1" + by (simp add: rat_number_expand) -lemma Fract_1_number_of: - "Fract 1 (number_of k) = 1 / number_of k" - unfolding Fract_of_int_quotient number_of_eq by simp subsubsection {* The ordered field of rational numbers *} @@ -771,7 +753,8 @@ (* not needed because x < (y::int) can be rewritten as x + 1 <= y: of_int_less_iff RS iffD2 *) #> Lin_Arith.add_simps [@{thm neg_less_iff_less}, @{thm True_implies_equals}, - read_instantiate @{context} [(("a", 0), "(number_of ?v)")] @{thm right_distrib}, + read_instantiate @{context} [(("a", 0), "(numeral ?v)")] @{thm right_distrib}, + read_instantiate @{context} [(("a", 0), "(neg_numeral ?v)")] @{thm right_distrib}, @{thm divide_1}, @{thm divide_zero_left}, @{thm times_divide_eq_right}, @{thm times_divide_eq_left}, @{thm minus_divide_left} RS sym, @{thm minus_divide_right} RS sym, @@ -895,9 +878,13 @@ lemma of_rat_of_int_eq [simp]: "of_rat (of_int z) = of_int z" by (cases z rule: int_diff_cases) (simp add: of_rat_diff) -lemma of_rat_number_of_eq [simp]: - "of_rat (number_of w) = (number_of w :: 'a::{number_ring,field_char_0})" -by (simp add: number_of_eq) +lemma of_rat_numeral_eq [simp]: + "of_rat (numeral w) = numeral w" +using of_rat_of_int_eq [of "numeral w"] by simp + +lemma of_rat_neg_numeral_eq [simp]: + "of_rat (neg_numeral w) = neg_numeral w" +using of_rat_of_int_eq [of "neg_numeral w"] by simp lemmas zero_rat = Zero_rat_def lemmas one_rat = One_rat_def @@ -935,9 +922,11 @@ lemma Rats_of_nat [simp]: "of_nat n \ Rats" by (subst of_rat_of_nat_eq [symmetric], rule Rats_of_rat) -lemma Rats_number_of [simp]: - "(number_of w::'a::{number_ring,field_char_0}) \ Rats" -by (subst of_rat_number_of_eq [symmetric], rule Rats_of_rat) +lemma Rats_number_of [simp]: "numeral w \ Rats" +by (subst of_rat_numeral_eq [symmetric], rule Rats_of_rat) + +lemma Rats_neg_number_of [simp]: "neg_numeral w \ Rats" +by (subst of_rat_neg_numeral_eq [symmetric], rule Rats_of_rat) lemma Rats_0 [simp]: "0 \ Rats" apply (unfold Rats_def) @@ -1032,6 +1021,8 @@ subsection {* Implementation of rational numbers as pairs of integers *} +text {* Formal constructor *} + definition Frct :: "int \ int \ rat" where [simp]: "Frct p = Fract (fst p) (snd p)" @@ -1039,17 +1030,45 @@ "Frct (quotient_of q) = q" by (cases q) (auto intro: quotient_of_eq) -lemma Frct_code_post [code_post]: - "Frct (0, k) = 0" - "Frct (k, 0) = 0" - "Frct (1, 1) = 1" - "Frct (number_of k, 1) = number_of k" - "Frct (1, number_of k) = 1 / number_of k" - "Frct (number_of k, number_of l) = number_of k / number_of l" - by (simp_all add: rat_number_collapse Fract_number_of_quotient Fract_1_number_of) + +text {* Numerals *} declare quotient_of_Fract [code abstract] +definition of_int :: "int \ rat" +where + [code_abbrev]: "of_int = Int.of_int" +hide_const (open) of_int + +lemma quotient_of_int [code abstract]: + "quotient_of (Rat.of_int a) = (a, 1)" + by (simp add: of_int_def of_int_rat quotient_of_Fract) + +lemma [code_unfold]: + "numeral k = Rat.of_int (numeral k)" + by (simp add: Rat.of_int_def) + +lemma [code_unfold]: + "neg_numeral k = Rat.of_int (neg_numeral k)" + by (simp add: Rat.of_int_def) + +lemma Frct_code_post [code_post]: + "Frct (0, a) = 0" + "Frct (a, 0) = 0" + "Frct (1, 1) = 1" + "Frct (numeral k, 1) = numeral k" + "Frct (neg_numeral k, 1) = neg_numeral k" + "Frct (1, numeral k) = 1 / numeral k" + "Frct (1, neg_numeral k) = 1 / neg_numeral k" + "Frct (numeral k, numeral l) = numeral k / numeral l" + "Frct (numeral k, neg_numeral l) = numeral k / neg_numeral l" + "Frct (neg_numeral k, numeral l) = neg_numeral k / numeral l" + "Frct (neg_numeral k, neg_numeral l) = neg_numeral k / neg_numeral l" + by (simp_all add: Fract_of_int_quotient) + + +text {* Operations *} + lemma rat_zero_code [code abstract]: "quotient_of 0 = (0, 1)" by (simp add: Zero_rat_def quotient_of_Fract normalize_def) @@ -1132,6 +1151,9 @@ "of_rat p = (let (a, b) = quotient_of p in of_int a / of_int b)" by (cases p) (simp add: quotient_of_Fract of_rat_rat) + +text {* Quickcheck *} + definition (in term_syntax) valterm_fract :: "int \ (unit \ Code_Evaluation.term) \ int \ (unit \ Code_Evaluation.term) \ rat \ (unit \ Code_Evaluation.term)" where [code_unfold]: "valterm_fract k l = Code_Evaluation.valtermify Fract {\} k {\} l" @@ -1212,7 +1234,6 @@ (@{const_name plus_rat_inst.plus_rat}, @{const_name Nitpick.plus_frac}), (@{const_name times_rat_inst.times_rat}, @{const_name Nitpick.times_frac}), (@{const_name uminus_rat_inst.uminus_rat}, @{const_name Nitpick.uminus_frac}), - (@{const_name number_rat_inst.number_of_rat}, @{const_name Nitpick.number_of_frac}), (@{const_name inverse_rat_inst.inverse_rat}, @{const_name Nitpick.inverse_frac}), (@{const_name ord_rat_inst.less_rat}, @{const_name Nitpick.less_frac}), (@{const_name ord_rat_inst.less_eq_rat}, @{const_name Nitpick.less_eq_frac}), @@ -1220,7 +1241,7 @@ *} lemmas [nitpick_unfold] = inverse_rat_inst.inverse_rat - number_rat_inst.number_of_rat one_rat_inst.one_rat ord_rat_inst.less_rat + one_rat_inst.one_rat ord_rat_inst.less_rat ord_rat_inst.less_eq_rat plus_rat_inst.plus_rat times_rat_inst.times_rat uminus_rat_inst.uminus_rat zero_rat_inst.zero_rat diff -r 29e92b644d6c -r f067afe98049 src/HOL/RealDef.thy --- a/src/HOL/RealDef.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/RealDef.thy Mon Mar 26 10:56:56 2012 +0200 @@ -720,7 +720,9 @@ unfolding less_eq_real_def less_real_def by (auto, drule (1) positive_add, simp add: positive_zero) show "a \ b \ c + a \ c + b" - unfolding less_eq_real_def less_real_def by auto + unfolding less_eq_real_def less_real_def by (auto simp: diff_minus) (* by auto *) + (* FIXME: Procedure int_combine_numerals: c + b - (c + a) \ b + - a *) + (* Should produce c + b - (c + a) \ b - a *) show "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)" by (rule sgn_real_def) show "a \ b \ b \ a" @@ -747,17 +749,6 @@ end -instantiation real :: number_ring -begin - -definition - "(number_of x :: real) = of_int x" - -instance proof -qed (rule number_of_real_def) - -end - lemma of_nat_Real: "of_nat x = Real (\n. of_nat x)" apply (induct x) apply (simp add: zero_real_def) @@ -877,7 +868,7 @@ by (erule contrapos_pp, simp add: not_less, erule Real_leI [OF Y]) lemma of_nat_less_two_power: - "of_nat n < (2::'a::{linordered_idom,number_ring}) ^ n" + "of_nat n < (2::'a::linordered_idom) ^ n" apply (induct n) apply simp apply (subgoal_tac "(1::'a) \ 2 ^ n") @@ -1469,18 +1460,19 @@ subsection{*Numerals and Arithmetic*} lemma [code_abbrev]: - "real_of_int (number_of k) = number_of k" - unfolding number_of_is_id number_of_real_def .. + "real_of_int (numeral k) = numeral k" + "real_of_int (neg_numeral k) = neg_numeral k" + by simp_all text{*Collapse applications of @{term real} to @{term number_of}*} -lemma real_number_of [simp]: "real (number_of v :: int) = number_of v" -by (simp add: real_of_int_def) +lemma real_numeral [simp]: + "real (numeral v :: int) = numeral v" + "real (neg_numeral v :: int) = neg_numeral v" +by (simp_all add: real_of_int_def) -lemma real_of_nat_number_of [simp]: - "real (number_of v :: nat) = - (if neg (number_of v :: int) then 0 - else (number_of v :: real))" -by (simp add: real_of_int_of_nat_eq [symmetric]) +lemma real_of_nat_numeral [simp]: + "real (numeral v :: nat) = numeral v" +by (simp add: real_of_nat_def) declaration {* K (Lin_Arith.add_inj_thms [@{thm real_of_nat_le_iff} RS iffD2, @{thm real_of_nat_inject} RS iffD2] @@ -1491,7 +1483,7 @@ @{thm real_of_nat_mult}, @{thm real_of_int_zero}, @{thm real_of_one}, @{thm real_of_int_add}, @{thm real_of_int_minus}, @{thm real_of_int_diff}, @{thm real_of_int_mult}, @{thm real_of_int_of_nat_eq}, - @{thm real_of_nat_number_of}, @{thm real_number_of}] + @{thm real_of_nat_numeral}, @{thm real_numeral(1)}, @{thm real_numeral(2)}] #> Lin_Arith.add_inj_const (@{const_name real}, @{typ "nat \ real"}) #> Lin_Arith.add_inj_const (@{const_name real}, @{typ "int \ real"})) *} @@ -1605,37 +1597,61 @@ subsection {* Implementation of rational real numbers *} +text {* Formal constructor *} + definition Ratreal :: "rat \ real" where - [simp]: "Ratreal = of_rat" + [code_abbrev, simp]: "Ratreal = of_rat" code_datatype Ratreal -lemma Ratreal_number_collapse [code_post]: - "Ratreal 0 = 0" - "Ratreal 1 = 1" - "Ratreal (number_of k) = number_of k" -by simp_all + +text {* Numerals *} + +lemma [code_abbrev]: + "(of_rat (of_int a) :: real) = of_int a" + by simp + +lemma [code_abbrev]: + "(of_rat 0 :: real) = 0" + by simp + +lemma [code_abbrev]: + "(of_rat 1 :: real) = 1" + by simp + +lemma [code_abbrev]: + "(of_rat (numeral k) :: real) = numeral k" + by simp -lemma zero_real_code [code, code_unfold]: +lemma [code_abbrev]: + "(of_rat (neg_numeral k) :: real) = neg_numeral k" + by simp + +lemma [code_post]: + "(of_rat (0 / r) :: real) = 0" + "(of_rat (r / 0) :: real) = 0" + "(of_rat (1 / 1) :: real) = 1" + "(of_rat (numeral k / 1) :: real) = numeral k" + "(of_rat (neg_numeral k / 1) :: real) = neg_numeral k" + "(of_rat (1 / numeral k) :: real) = 1 / numeral k" + "(of_rat (1 / neg_numeral k) :: real) = 1 / neg_numeral k" + "(of_rat (numeral k / numeral l) :: real) = numeral k / numeral l" + "(of_rat (numeral k / neg_numeral l) :: real) = numeral k / neg_numeral l" + "(of_rat (neg_numeral k / numeral l) :: real) = neg_numeral k / numeral l" + "(of_rat (neg_numeral k / neg_numeral l) :: real) = neg_numeral k / neg_numeral l" + by (simp_all add: of_rat_divide) + + +text {* Operations *} + +lemma zero_real_code [code]: "0 = Ratreal 0" by simp -lemma one_real_code [code, code_unfold]: +lemma one_real_code [code]: "1 = Ratreal 1" by simp -lemma number_of_real_code [code_unfold]: - "number_of k = Ratreal (number_of k)" -by simp - -lemma Ratreal_number_of_quotient [code_post]: - "Ratreal (number_of r) / Ratreal (number_of s) = number_of r / number_of s" -by simp - -lemma Ratreal_number_of_quotient2 [code_post]: - "Ratreal (number_of r / number_of s) = number_of r / number_of s" -unfolding Ratreal_number_of_quotient [symmetric] Ratreal_def of_rat_divide .. - instantiation real :: equal begin @@ -1681,6 +1697,9 @@ lemma real_floor_code [code]: "floor (Ratreal x) = floor x" by (metis Ratreal_def floor_le_iff floor_unique le_floor_iff of_int_floor_le of_rat_of_int_eq real_less_eq_code) + +text {* Quickcheck *} + definition (in term_syntax) valterm_ratreal :: "rat \ (unit \ Code_Evaluation.term) \ real \ (unit \ Code_Evaluation.term)" where [code_unfold]: "valterm_ratreal k = Code_Evaluation.valtermify Ratreal {\} k" @@ -1741,14 +1760,12 @@ (@{const_name plus_real_inst.plus_real}, @{const_name Nitpick.plus_frac}), (@{const_name times_real_inst.times_real}, @{const_name Nitpick.times_frac}), (@{const_name uminus_real_inst.uminus_real}, @{const_name Nitpick.uminus_frac}), - (@{const_name number_real_inst.number_of_real}, @{const_name Nitpick.number_of_frac}), (@{const_name inverse_real_inst.inverse_real}, @{const_name Nitpick.inverse_frac}), (@{const_name ord_real_inst.less_real}, @{const_name Nitpick.less_frac}), (@{const_name ord_real_inst.less_eq_real}, @{const_name Nitpick.less_eq_frac})] *} -lemmas [nitpick_unfold] = inverse_real_inst.inverse_real - number_real_inst.number_of_real one_real_inst.one_real +lemmas [nitpick_unfold] = inverse_real_inst.inverse_real one_real_inst.one_real ord_real_inst.less_real ord_real_inst.less_eq_real plus_real_inst.plus_real times_real_inst.times_real uminus_real_inst.uminus_real zero_real_inst.zero_real diff -r 29e92b644d6c -r f067afe98049 src/HOL/RealVector.thy --- a/src/HOL/RealVector.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/RealVector.thy Mon Mar 26 10:56:56 2012 +0200 @@ -303,9 +303,11 @@ lemma of_real_of_int_eq [simp]: "of_real (of_int z) = of_int z" by (cases z rule: int_diff_cases, simp) -lemma of_real_number_of_eq: - "of_real (number_of w) = (number_of w :: 'a::{number_ring,real_algebra_1})" -by (simp add: number_of_eq) +lemma of_real_numeral: "of_real (numeral w) = numeral w" +using of_real_of_int_eq [of "numeral w"] by simp + +lemma of_real_neg_numeral: "of_real (neg_numeral w) = neg_numeral w" +using of_real_of_int_eq [of "neg_numeral w"] by simp text{*Every real algebra has characteristic zero*} @@ -335,9 +337,11 @@ lemma Reals_of_nat [simp]: "of_nat n \ Reals" by (subst of_real_of_nat_eq [symmetric], rule Reals_of_real) -lemma Reals_number_of [simp]: - "(number_of w::'a::{number_ring,real_algebra_1}) \ Reals" -by (subst of_real_number_of_eq [symmetric], rule Reals_of_real) +lemma Reals_numeral [simp]: "numeral w \ Reals" +by (subst of_real_numeral [symmetric], rule Reals_of_real) + +lemma Reals_neg_numeral [simp]: "neg_numeral w \ Reals" +by (subst of_real_neg_numeral [symmetric], rule Reals_of_real) lemma Reals_0 [simp]: "0 \ Reals" apply (unfold Reals_def) @@ -752,10 +756,13 @@ "norm (of_real r :: 'a::real_normed_algebra_1) = \r\" unfolding of_real_def by simp -lemma norm_number_of [simp]: - "norm (number_of w::'a::{number_ring,real_normed_algebra_1}) - = \number_of w\" -by (subst of_real_number_of_eq [symmetric], rule norm_of_real) +lemma norm_numeral [simp]: + "norm (numeral w::'a::real_normed_algebra_1) = numeral w" +by (subst of_real_numeral [symmetric], subst norm_of_real, simp) + +lemma norm_neg_numeral [simp]: + "norm (neg_numeral w::'a::real_normed_algebra_1) = numeral w" +by (subst of_real_neg_numeral [symmetric], subst norm_of_real, simp) lemma norm_of_int [simp]: "norm (of_int z::'a::real_normed_algebra_1) = \of_int z\" diff -r 29e92b644d6c -r f067afe98049 src/HOL/SMT_Examples/SMT_Examples.thy --- a/src/HOL/SMT_Examples/SMT_Examples.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/SMT_Examples/SMT_Examples.thy Mon Mar 26 10:56:56 2012 +0200 @@ -467,7 +467,7 @@ lemma "(f g (x::'a::type) = (g x \ True)) \ (f g x = True) \ (g x = True)" by smt -lemma "id 3 = 3 \ id True = True" by (smt id_def) +lemma "id x = x \ id True = True" (* BROKEN by (smt id_def) *) oops lemma "i \ i1 \ i \ i2 \ ((f (i1 := v1)) (i2 := v2)) i = f i" using fun_upd_same fun_upd_apply diff -r 29e92b644d6c -r f067afe98049 src/HOL/SMT_Examples/SMT_Tests.thy --- a/src/HOL/SMT_Examples/SMT_Tests.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/SMT_Examples/SMT_Tests.thy Mon Mar 26 10:56:56 2012 +0200 @@ -211,8 +211,8 @@ lemma assumes "\x y. SMT.trigger [[SMT.pat (f x), SMT.pat (g y)]] (f x = g y)" - shows "f 1 = g 2" - using assms by smt + shows "f a = g b" + using assms (* BROKEN by smt *) oops lemma assumes "ALL x. SMT.trigger [[SMT.pat (P x)]] (P x --> Q x)" @@ -853,7 +853,7 @@ using point.simps using [[smt_datatypes, smt_oracle]] using [[z3_options="AUTO_CONFIG=false"]] - by smt+ + (* BROKEN by smt+ *) oops lemma "cy (p \ cx := a \) = cy p" @@ -862,7 +862,7 @@ using point.simps using [[smt_datatypes, smt_oracle]] using [[z3_options="AUTO_CONFIG=false"]] - by smt+ + (* BROKEN by smt+ *) oops lemma "p1 = p2 \ cx p1 = cx p2" @@ -891,7 +891,7 @@ using point.simps bw_point.simps using [[smt_datatypes, smt_oracle]] using [[z3_options="AUTO_CONFIG=false"]] - by smt+ + (* BROKEN by smt+ *) oops lemma "\ cx = 3, cy = 4, black = b \ \ black := w \ = \ cx = 3, cy = 4, black = w \" @@ -905,7 +905,7 @@ using point.simps bw_point.simps using [[smt_datatypes, smt_oracle]] using [[z3_options="AUTO_CONFIG=false"]] - by smt + (* BROKEN by smt *) oops subsubsection {* Type definitions *} @@ -919,7 +919,7 @@ using n1_def n2_def n3_def nplus_def using [[smt_datatypes, smt_oracle]] using [[z3_options="AUTO_CONFIG=false"]] - by smt+ + (* BROKEN by smt+ *) oops diff -r 29e92b644d6c -r f067afe98049 src/HOL/SPARK/SPARK.thy --- a/src/HOL/SPARK/SPARK.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/SPARK/SPARK.thy Mon Mar 26 10:56:56 2012 +0200 @@ -145,7 +145,7 @@ then have "bin = 0" "bit = 0" by (auto simp add: Bit_def bitval_def split add: bit.split_asm) arith then show ?thesis using 0 1 `y < 2 ^ n` - by simp (simp add: Bit0_def int_or_Pls [unfolded Pls_def]) + by simp next case (Suc m) from 3 have "0 \ bin" @@ -188,7 +188,7 @@ then have "bin = 0" "bit = 0" by (auto simp add: Bit_def bitval_def split add: bit.split_asm) arith then show ?thesis using 0 1 `y < 2 ^ n` - by simp (simp add: Bit0_def int_xor_Pls [unfolded Pls_def]) + by simp next case (Suc m) from 3 have "0 \ bin" @@ -257,17 +257,17 @@ proof (induct x arbitrary: n rule: bin_induct) case 1 then show ?case - by simp (simp add: Pls_def) + by simp next case 2 then show ?case - by (simp, simp only: Min_def, simp add: m1mod2k) + by (simp, simp add: m1mod2k) next case (3 bin bit) show ?case proof (cases n) case 0 - then show ?thesis by (simp add: int_and_extra_simps [unfolded Pls_def]) + then show ?thesis by (simp add: int_and_extra_simps) next case (Suc m) with 3 show ?thesis diff -r 29e92b644d6c -r f067afe98049 src/HOL/Semiring_Normalization.thy --- a/src/HOL/Semiring_Normalization.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Semiring_Normalization.thy Mon Mar 26 10:56:56 2012 +0200 @@ -116,7 +116,8 @@ "x ^ (Suc q) = x * (x ^ q)" "x ^ (2*n) = (x ^ n) * (x ^ n)" "x ^ (Suc (2*n)) = x * ((x ^ n) * (x ^ n))" - by (simp_all add: algebra_simps power_add power2_eq_square power_mult_distrib power_mult) + by (simp_all add: algebra_simps power_add power2_eq_square + power_mult_distrib power_mult del: one_add_one) lemmas normalizing_comm_semiring_1_axioms = comm_semiring_1_axioms [normalizer @@ -218,4 +219,13 @@ hide_fact (open) normalizing_field_axioms normalizing_field_ops normalizing_field_rules +code_modulename SML + Semiring_Normalization Arith + +code_modulename OCaml + Semiring_Normalization Arith + +code_modulename Haskell + Semiring_Normalization Arith + end diff -r 29e92b644d6c -r f067afe98049 src/HOL/Series.thy --- a/src/HOL/Series.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Series.thy Mon Mar 26 10:56:56 2012 +0200 @@ -417,8 +417,8 @@ shows "norm x < 1 \ summable (\n. x ^ n)" by (rule geometric_sums [THEN sums_summable]) -lemma half: "0 < 1 / (2::'a::{number_ring,linordered_field_inverse_zero})" - by arith +lemma half: "0 < 1 / (2::'a::linordered_field)" + by simp lemma power_half_series: "(\n. (1/2::real)^Suc n) sums 1" proof - diff -r 29e92b644d6c -r f067afe98049 src/HOL/SetInterval.thy --- a/src/HOL/SetInterval.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/SetInterval.thy Mon Mar 26 10:56:56 2012 +0200 @@ -1282,7 +1282,7 @@ subsection {* The formula for arithmetic sums *} -lemma gauss_sum: +lemma gauss_sum: (* FIXME: rephrase in terms of "2" *) "((1::'a::comm_semiring_1) + 1)*(\i\{1..n}. of_nat i) = of_nat n*((of_nat n)+1)" proof (induct n) @@ -1290,7 +1290,7 @@ show ?case by simp next case (Suc n) - then show ?case by (simp add: algebra_simps) + then show ?case by (simp add: algebra_simps del: one_add_one) (* FIXME *) qed theorem arith_series_general: @@ -1308,18 +1308,18 @@ unfolding One_nat_def by (simp add: setsum_right_distrib atLeast0LessThan[symmetric] setsum_shift_lb_Suc0_0_upt mult_ac) also have "(1+1)*\ = (1+1)*?n*a + d*(1+1)*(\i\{1..i\{1..n - 1}. ?I i) = ((1+1)*?n*a + d*?I (n - 1)*?I n)" by (simp only: mult_ac gauss_sum [of "n - 1"], unfold One_nat_def) (simp add: mult_ac trans [OF add_commute of_nat_Suc [symmetric]]) - finally show ?thesis by (simp add: algebra_simps) + finally show ?thesis by (simp add: algebra_simps del: one_add_one) next assume "\(n > 1)" hence "n = 1 \ n = 0" by auto - thus ?thesis by (auto simp: algebra_simps) + thus ?thesis by (auto simp: algebra_simps mult_2_right) qed lemma arith_series_nat: diff -r 29e92b644d6c -r f067afe98049 src/HOL/Tools/Nitpick/nitpick.ML --- a/src/HOL/Tools/Nitpick/nitpick.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Tools/Nitpick/nitpick.ML Mon Mar 26 10:56:56 2012 +0200 @@ -206,7 +206,7 @@ val syntactic_sorts = @{sort "{default,zero,one,plus,minus,uminus,times,inverse,abs,sgn,ord,equal}"} @ - @{sort number} + @{sort numeral} fun has_tfree_syntactic_sort (TFree (_, S as _ :: _)) = subset (op =) (S, syntactic_sorts) | has_tfree_syntactic_sort _ = false diff -r 29e92b644d6c -r f067afe98049 src/HOL/Tools/Nitpick/nitpick_hol.ML --- a/src/HOL/Tools/Nitpick/nitpick_hol.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML Mon Mar 26 10:56:56 2012 +0200 @@ -1636,30 +1636,32 @@ (hol_ctxt as {thy, ctxt, stds, whacks, total_consts, case_names, def_tables, ground_thm_table, ersatz_table, ...}) = let - fun do_term depth Ts t = + fun do_numeral depth Ts mult T t0 t1 = + (if is_number_type ctxt T then + let + val j = mult * (HOLogic.dest_num t1) + |> T = nat_T ? Integer.max 0 + val s = numeral_prefix ^ signed_string_of_int j + in + if is_integer_like_type T then + if is_standard_datatype thy stds T then Const (s, T) + else funpow j (curry (op $) (suc_const T)) (zero_const T) + else + do_term depth Ts (Const (@{const_name of_int}, int_T --> T) + $ Const (s, int_T)) + end + handle TERM _ => raise SAME () + else + raise SAME ()) + handle SAME () => s_betapply [] (do_term depth Ts t0, do_term depth Ts t1) + and do_term depth Ts t = case t of - (t0 as Const (@{const_name Int.number_class.number_of}, + (t0 as Const (@{const_name Num.neg_numeral_class.neg_numeral}, Type (@{type_name fun}, [_, ran_T]))) $ t1 => - ((if is_number_type ctxt ran_T then - let - val j = t1 |> HOLogic.dest_numeral - |> ran_T = nat_T ? Integer.max 0 - val s = numeral_prefix ^ signed_string_of_int j - in - if is_integer_like_type ran_T then - if is_standard_datatype thy stds ran_T then - Const (s, ran_T) - else - funpow j (curry (op $) (suc_const ran_T)) (zero_const ran_T) - else - do_term depth Ts (Const (@{const_name of_int}, int_T --> ran_T) - $ Const (s, int_T)) - end - handle TERM _ => raise SAME () - else - raise SAME ()) - handle SAME () => - s_betapply [] (do_term depth Ts t0, do_term depth Ts t1)) + do_numeral depth Ts ~1 ran_T t0 t1 + | (t0 as Const (@{const_name Num.numeral_class.numeral}, + Type (@{type_name fun}, [_, ran_T]))) $ t1 => + do_numeral depth Ts 1 ran_T t0 t1 | Const (@{const_name refl_on}, T) $ Const (@{const_name top}, _) $ t2 => do_const depth Ts t (@{const_name refl'}, range_type T) [t2] | (t0 as Const (@{const_name Sigma}, Type (_, [T1, Type (_, [T2, T3])]))) diff -r 29e92b644d6c -r f067afe98049 src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML Mon Mar 26 10:56:56 2012 +0200 @@ -240,10 +240,12 @@ @{const_name Groups.one}, @{const_name Groups.plus}, @{const_name Nat.ord_nat_inst.less_eq_nat}, @{const_name Nat.ord_nat_inst.less_nat}, +(* FIXME @{const_name number_nat_inst.number_of_nat}, - @{const_name Int.Bit0}, - @{const_name Int.Bit1}, - @{const_name Int.Pls}, +*) + @{const_name Num.Bit0}, + @{const_name Num.Bit1}, + @{const_name Num.One}, @{const_name Int.zero_int_inst.zero_int}, @{const_name List.filter}, @{const_name HOL.If}, diff -r 29e92b644d6c -r f067afe98049 src/HOL/Tools/Qelim/cooper.ML --- a/src/HOL/Tools/Qelim/cooper.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Tools/Qelim/cooper.ML Mon Mar 26 10:56:56 2012 +0200 @@ -41,9 +41,9 @@ @{term "Ex :: (int => _) => _"}, @{term "Ex :: (nat => _) => _"}, @{term "All :: (int => _) => _"}, @{term "All :: (nat => _) => _"}, @{term "nat"}, @{term "int"}, - @{term "Int.Bit0"}, @{term "Int.Bit1"}, - @{term "Int.Pls"}, @{term "Int.Min"}, - @{term "Int.number_of :: int => int"}, @{term "Int.number_of :: int => nat"}, + @{term "Num.One"}, @{term "Num.Bit0"}, @{term "Num.Bit1"}, + @{term "Num.numeral :: num => int"}, @{term "Num.numeral :: num => nat"}, + @{term "Num.neg_numeral :: num => int"}, @{term "0::int"}, @{term "1::int"}, @{term "0::nat"}, @{term "1::nat"}, @{term "True"}, @{term "False"}]; @@ -595,8 +595,10 @@ | num_of_term vs (Term.Bound i) = Proc.Bound i | num_of_term vs @{term "0::int"} = Proc.C 0 | num_of_term vs @{term "1::int"} = Proc.C 1 - | num_of_term vs (t as Const (@{const_name number_of}, _) $ _) = + | num_of_term vs (t as Const (@{const_name numeral}, _) $ _) = Proc.C (dest_number t) + | num_of_term vs (t as Const (@{const_name neg_numeral}, _) $ _) = + Proc.Neg (Proc.C (dest_number t)) | num_of_term vs (Const (@{const_name Groups.uminus}, _) $ t') = Proc.Neg (num_of_term vs t') | num_of_term vs (Const (@{const_name Groups.plus}, _) $ t1 $ t2) = @@ -784,16 +786,16 @@ local val ss1 = comp_ss - addsimps @{thms simp_thms} @ [@{thm "nat_number_of_def"}, @{thm "zdvd_int"}] - @ map (fn r => r RS sym) + addsimps @{thms simp_thms} @ [@{thm "nat_numeral"} RS sym, @{thm "zdvd_int"}] + @ map (fn r => r RS sym) [@{thm "int_int_eq"}, @{thm "zle_int"}, @{thm "zless_int"}, @{thm "zadd_int"}, @{thm "zmult_int"}] |> Splitter.add_split @{thm "zdiff_int_split"} val ss2 = HOL_basic_ss - addsimps [@{thm "nat_0_le"}, @{thm "int_nat_number_of"}, - @{thm "all_nat"}, @{thm "ex_nat"}, @{thm "number_of1"}, - @{thm "number_of2"}, @{thm "int_0"}, @{thm "int_1"}, @{thm "Suc_eq_plus1"}] + addsimps [@{thm "nat_0_le"}, @{thm "int_numeral"}, + @{thm "all_nat"}, @{thm "ex_nat"}, @{thm "zero_le_numeral"}, + @{thm "le_numeral_extra"(3)}, @{thm "int_0"}, @{thm "int_1"}, @{thm "Suc_eq_plus1"}] |> fold Simplifier.add_cong [@{thm "conj_le_cong"}, @{thm "imp_le_cong"}] val div_mod_ss = HOL_basic_ss addsimps @{thms simp_thms} @ map (Thm.symmetric o mk_meta_eq) diff -r 29e92b644d6c -r f067afe98049 src/HOL/Tools/Quickcheck/PNF_Narrowing_Engine.hs --- a/src/HOL/Tools/Quickcheck/PNF_Narrowing_Engine.hs Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Tools/Quickcheck/PNF_Narrowing_Engine.hs Mon Mar 26 10:56:56 2012 +0200 @@ -301,7 +301,7 @@ checkOf (Generated_Code.Universal _ f _) = (\(t : ts) -> checkOf (f t) ts) checkOf (Generated_Code.Existential _ f _) = (\(t : ts) -> checkOf (f t) ts) -dummy = Generated_Code.Var [] (Generated_Code.SumOfProd [[]]) +dummy = Generated_Code.Narrowing_variable [] (Generated_Code.Narrowing_sum_of_products [[]]) treeOf :: Int -> Generated_Code.Property -> QuantTree treeOf n (Generated_Code.Property _) = Node uneval diff -r 29e92b644d6c -r f067afe98049 src/HOL/Tools/Quotient/quotient_def.ML --- a/src/HOL/Tools/Quotient/quotient_def.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Tools/Quotient/quotient_def.ML Mon Mar 26 10:56:56 2012 +0200 @@ -6,13 +6,17 @@ signature QUOTIENT_DEF = sig + val add_quotient_def: + ((binding * mixfix) * Attrib.binding) * (term * term) -> thm -> + local_theory -> Quotient_Info.quotconsts * local_theory + val quotient_def: (binding * typ option * mixfix) option * (Attrib.binding * (term * term)) -> - local_theory -> Quotient_Info.quotconsts * local_theory + local_theory -> Proof.state val quotient_def_cmd: (binding * string option * mixfix) option * (Attrib.binding * (string * string)) -> - local_theory -> Quotient_Info.quotconsts * local_theory + local_theory -> Proof.state val lift_raw_const: typ list -> (string * term * mixfix) -> local_theory -> Quotient_Info.quotconsts * local_theory @@ -23,6 +27,130 @@ (** Interface and Syntax Setup **) +(* Generation of the code certificate from the rsp theorem *) + +infix 0 MRSL + +fun ants MRSL thm = fold (fn rl => fn thm => rl RS thm) ants thm + +fun get_body_types (Type ("fun", [_, U]), Type ("fun", [_, V])) = get_body_types (U, V) + | get_body_types (U, V) = (U, V) + +fun get_binder_types (Type ("fun", [T, U]), Type ("fun", [V, W])) = (T, V) :: get_binder_types (U, W) + | get_binder_types _ = [] + +fun unabs_def ctxt def = + let + val (_, rhs) = Thm.dest_equals (cprop_of def) + fun dest_abs (Abs (var_name, T, _)) = (var_name, T) + | dest_abs tm = raise TERM("get_abs_var",[tm]) + val (var_name, T) = dest_abs (term_of rhs) + val (new_var_names, ctxt') = Variable.variant_fixes [var_name] ctxt + val thy = Proof_Context.theory_of ctxt' + val refl_thm = Thm.reflexive (cterm_of thy (Free (hd new_var_names, T))) + in + Thm.combination def refl_thm |> + singleton (Proof_Context.export ctxt' ctxt) + end + +fun unabs_all_def ctxt def = + let + val (_, rhs) = Thm.dest_equals (cprop_of def) + val xs = strip_abs_vars (term_of rhs) + in + fold (K (unabs_def ctxt)) xs def + end + +val map_fun_unfolded = + @{thm map_fun_def[abs_def]} |> + unabs_def @{context} |> + unabs_def @{context} |> + Local_Defs.unfold @{context} [@{thm comp_def}] + +fun unfold_fun_maps ctm = + let + fun unfold_conv ctm = + case (Thm.term_of ctm) of + Const (@{const_name "map_fun"}, _) $ _ $ _ => + (Conv.arg_conv unfold_conv then_conv Conv.rewr_conv map_fun_unfolded) ctm + | _ => Conv.all_conv ctm + val try_beta_conv = Conv.try_conv (Thm.beta_conversion false) + in + (Conv.arg_conv (Conv.fun_conv unfold_conv then_conv try_beta_conv)) ctm + end + +fun prove_rel ctxt rsp_thm (rty, qty) = + let + val ty_args = get_binder_types (rty, qty) + fun disch_arg args_ty thm = + let + val quot_thm = Quotient_Term.prove_quot_theorem ctxt args_ty + in + [quot_thm, thm] MRSL @{thm apply_rsp''} + end + in + fold disch_arg ty_args rsp_thm + end + +exception CODE_CERT_GEN of string + +fun simplify_code_eq ctxt def_thm = + Local_Defs.unfold ctxt [@{thm o_def}, @{thm map_fun_def}, @{thm id_def}] def_thm + +fun generate_code_cert ctxt def_thm rsp_thm (rty, qty) = + let + val quot_thm = Quotient_Term.prove_quot_theorem ctxt (get_body_types (rty, qty)) + val fun_rel = prove_rel ctxt rsp_thm (rty, qty) + val abs_rep_thm = [quot_thm, fun_rel] MRSL @{thm Quotient_rep_abs} + val abs_rep_eq = + case (HOLogic.dest_Trueprop o prop_of) fun_rel of + Const (@{const_name HOL.eq}, _) $ _ $ _ => abs_rep_thm + | Const (@{const_name invariant}, _) $ _ $ _ $ _ => abs_rep_thm RS @{thm invariant_to_eq} + | _ => raise CODE_CERT_GEN "relation is neither equality nor invariant" + val unfolded_def = Conv.fconv_rule unfold_fun_maps def_thm + val unabs_def = unabs_all_def ctxt unfolded_def + val rep = (snd o Thm.dest_comb o snd o Thm.dest_comb o cprop_of) quot_thm + val rep_refl = Thm.reflexive rep RS @{thm meta_eq_to_obj_eq} + val repped_eq = [rep_refl, unabs_def RS @{thm meta_eq_to_obj_eq}] MRSL @{thm cong} + val code_cert = [repped_eq, abs_rep_eq] MRSL @{thm trans} + in + simplify_code_eq ctxt code_cert + end + +fun define_code_cert def_thm rsp_thm (rty, qty) lthy = + let + val quot_thm = Quotient_Term.prove_quot_theorem lthy (get_body_types (rty, qty)) + in + if Quotient_Type.can_generate_code_cert quot_thm then + let + val code_cert = generate_code_cert lthy def_thm rsp_thm (rty, qty) + val add_abs_eqn_attribute = + Thm.declaration_attribute (fn thm => Context.mapping (Code.add_abs_eqn thm) I) + val add_abs_eqn_attrib = Attrib.internal (K add_abs_eqn_attribute); + in + lthy + |> (snd oo Local_Theory.note) ((Binding.empty, [add_abs_eqn_attrib]), [code_cert]) + end + else + lthy + end + +fun define_code_eq def_thm lthy = + let + val unfolded_def = Conv.fconv_rule unfold_fun_maps def_thm + val code_eq = unabs_all_def lthy unfolded_def + val simp_code_eq = simplify_code_eq lthy code_eq + in + lthy + |> (snd oo Local_Theory.note) ((Binding.empty, [Code.add_default_eqn_attrib]), [simp_code_eq]) + end + +fun define_code def_thm rsp_thm (rty, qty) lthy = + if body_type rty = body_type qty then + define_code_eq def_thm lthy + else + define_code_cert def_thm rsp_thm (rty, qty) lthy + (* The ML-interface for a quotient definition takes as argument: @@ -30,6 +158,7 @@ - attributes - the new constant as term - the rhs of the definition as term + - respectfulness theorem for the rhs It stores the qconst_info in the quotconsts data slot. @@ -45,7 +174,84 @@ quote str ^ " differs from declaration " ^ name ^ pos) end -fun gen_quotient_def prep_vars prep_term (raw_var, ((name, atts), (lhs_raw, rhs_raw))) lthy = +fun add_quotient_def ((var, (name, atts)), (lhs, rhs)) rsp_thm lthy = + let + val rty = fastype_of rhs + val qty = fastype_of lhs + val absrep_trm = + Quotient_Term.absrep_fun lthy Quotient_Term.AbsF (rty, qty) $ rhs + val prop = Syntax.check_term lthy (Logic.mk_equals (lhs, absrep_trm)) + val (_, prop') = Local_Defs.cert_def lthy prop + val (_, newrhs) = Local_Defs.abs_def prop' + + val ((trm, (_ , def_thm)), lthy') = + Local_Theory.define (var, ((Thm.def_binding_optional (#1 var) name, atts), newrhs)) lthy + + (* data storage *) + val qconst_data = {qconst = trm, rconst = rhs, def = def_thm} + fun get_rsp_thm_name (lhs_name, _) = Binding.suffix_name "_rsp" lhs_name + + val lthy'' = lthy' + |> Local_Theory.declaration {syntax = false, pervasive = true} + (fn phi => + (case Quotient_Info.transform_quotconsts phi qconst_data of + qcinfo as {qconst = Const (c, _), ...} => + Quotient_Info.update_quotconsts c qcinfo + | _ => I)) + |> (snd oo Local_Theory.note) + ((get_rsp_thm_name var, [Attrib.internal (K Quotient_Info.rsp_rules_add)]), + [rsp_thm]) + |> define_code def_thm rsp_thm (rty, qty) + + in + (qconst_data, lthy'') + end + +fun mk_readable_rsp_thm_eq tm lthy = + let + val ctm = cterm_of (Proof_Context.theory_of lthy) tm + + fun norm_fun_eq ctm = + let + fun abs_conv2 cv = Conv.abs_conv (K (Conv.abs_conv (K cv) lthy)) lthy + fun erase_quants ctm' = + case (Thm.term_of ctm') of + Const ("HOL.eq", _) $ _ $ _ => Conv.all_conv ctm' + | _ => (Conv.binder_conv (K erase_quants) lthy then_conv + Conv.rewr_conv @{thm fun_eq_iff[symmetric, THEN eq_reflection]}) ctm' + in + (abs_conv2 erase_quants then_conv Thm.eta_conversion) ctm + end + + fun simp_arrows_conv ctm = + let + val unfold_conv = Conv.rewrs_conv + [@{thm fun_rel_eq_invariant[THEN eq_reflection]}, @{thm fun_rel_eq_rel[THEN eq_reflection]}, + @{thm fun_rel_def[THEN eq_reflection]}] + val left_conv = simp_arrows_conv then_conv Conv.try_conv norm_fun_eq + fun binop_conv2 cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2 + in + case (Thm.term_of ctm) of + Const (@{const_name "fun_rel"}, _) $ _ $ _ => + (binop_conv2 left_conv simp_arrows_conv then_conv unfold_conv) ctm + | _ => Conv.all_conv ctm + end + + val unfold_ret_val_invs = Conv.bottom_conv + (K (Conv.try_conv (Conv.rewr_conv @{thm invariant_same_args}))) lthy + val simp_conv = Conv.arg_conv (Conv.fun2_conv simp_arrows_conv) + val univq_conv = Conv.rewr_conv @{thm HOL.all_simps(6)[symmetric, THEN eq_reflection]} + val univq_prenex_conv = Conv.top_conv (K (Conv.try_conv univq_conv)) lthy + val beta_conv = Thm.beta_conversion true + val eq_thm = + (simp_conv then_conv univq_prenex_conv then_conv beta_conv then_conv unfold_ret_val_invs) ctm + in + Object_Logic.rulify(eq_thm RS Drule.equal_elim_rule2) + end + + + +fun gen_quotient_def prep_vars prep_term (raw_var, (attr, (lhs_raw, rhs_raw))) lthy = let val (vars, ctxt) = prep_vars (the_list raw_var) lthy val T_opt = (case vars of [(_, SOME T, _)] => SOME T | _ => NONE) @@ -63,27 +269,50 @@ if Variable.check_name binding = lhs_str then (binding, mx) else error_msg binding lhs_str | _ => raise Match) - - val absrep_trm = Quotient_Term.absrep_fun lthy Quotient_Term.AbsF (fastype_of rhs, lhs_ty) $ rhs - val prop = Syntax.check_term lthy (Logic.mk_equals (lhs, absrep_trm)) - val (_, prop') = Local_Defs.cert_def lthy prop - val (_, newrhs) = Local_Defs.abs_def prop' - - val ((trm, (_ , thm)), lthy') = - Local_Theory.define (var, ((Thm.def_binding_optional (#1 var) name, atts), newrhs)) lthy + + fun try_to_prove_refl thm = + let + val lhs_eq = + thm + |> prop_of + |> Logic.dest_implies + |> fst + |> strip_all_body + |> try HOLogic.dest_Trueprop + in + case lhs_eq of + SOME (Const ("HOL.eq", _) $ _ $ _) => SOME (@{thm refl} RS thm) + | SOME _ => (case body_type (fastype_of lhs) of + Type (typ_name, _) => ( SOME + (#equiv_thm (the (Quotient_Info.lookup_quotients lthy typ_name)) + RS @{thm Equiv_Relations.equivp_reflp} RS thm) + handle _ => NONE) + | _ => NONE + ) + | _ => NONE + end - (* data storage *) - val qconst_data = {qconst = trm, rconst = rhs, def = thm} + val rsp_rel = Quotient_Term.equiv_relation lthy (fastype_of rhs, lhs_ty) + val internal_rsp_tm = HOLogic.mk_Trueprop (Syntax.check_term lthy (rsp_rel $ rhs $ rhs)) + val readable_rsp_thm_eq = mk_readable_rsp_thm_eq internal_rsp_tm lthy + val maybe_proven_rsp_thm = try_to_prove_refl readable_rsp_thm_eq + val (readable_rsp_tm, _) = Logic.dest_implies (prop_of readable_rsp_thm_eq) + + fun after_qed thm_list lthy = + let + val internal_rsp_thm = + case thm_list of + [] => the maybe_proven_rsp_thm + | [[thm]] => Goal.prove ctxt [] [] internal_rsp_tm + (fn _ => rtac readable_rsp_thm_eq 1 THEN Proof_Context.fact_tac [thm] 1) + in + snd (add_quotient_def ((var, attr), (lhs, rhs)) internal_rsp_thm lthy) + end - val lthy'' = lthy' - |> Local_Theory.declaration {syntax = false, pervasive = true} - (fn phi => - (case Quotient_Info.transform_quotconsts phi qconst_data of - qcinfo as {qconst = Const (c, _), ...} => - Quotient_Info.update_quotconsts c qcinfo - | _ => I)); in - (qconst_data, lthy'') + case maybe_proven_rsp_thm of + SOME _ => Proof.theorem NONE after_qed [] lthy + | NONE => Proof.theorem NONE after_qed [[(readable_rsp_tm,[])]] lthy end fun check_term' cnstr ctxt = @@ -103,16 +332,19 @@ val qty = Quotient_Term.derive_qtyp ctxt qtys rty val lhs = Free (qconst_name, qty) in - quotient_def (SOME (Binding.name qconst_name, NONE, mx), (Attrib.empty_binding, (lhs, rconst))) ctxt + (*quotient_def (SOME (Binding.name qconst_name, NONE, mx), (Attrib.empty_binding, (lhs, rconst))) ctxt*) + ({qconst = lhs, rconst = lhs, def = @{thm refl}}, ctxt) end -(* command *) +(* parser and command *) +val quotdef_parser = + Scan.option Parse_Spec.constdecl -- + Parse.!!! (Parse_Spec.opt_thm_name ":" -- (Parse.term --| @{keyword "is"} -- Parse.term)) val _ = - Outer_Syntax.local_theory @{command_spec "quotient_definition"} + Outer_Syntax.local_theory_to_proof @{command_spec "quotient_definition"} "definition for constants over the quotient type" - (Scan.option Parse_Spec.constdecl -- - Parse.!!! (Parse_Spec.opt_thm_name ":" -- (Parse.term --| @{keyword "is"} -- Parse.term)) - >> (snd oo quotient_def_cmd)) + (quotdef_parser >> quotient_def_cmd) + end; (* structure *) diff -r 29e92b644d6c -r f067afe98049 src/HOL/Tools/Quotient/quotient_info.ML --- a/src/HOL/Tools/Quotient/quotient_info.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Tools/Quotient/quotient_info.ML Mon Mar 26 10:56:56 2012 +0200 @@ -6,7 +6,7 @@ signature QUOTIENT_INFO = sig - type quotmaps = {relmap: string} + type quotmaps = {relmap: string, quot_thm: thm} val lookup_quotmaps: Proof.context -> string -> quotmaps option val lookup_quotmaps_global: theory -> string -> quotmaps option val print_quotmaps: Proof.context -> unit @@ -18,7 +18,7 @@ val update_abs_rep: string -> abs_rep -> Context.generic -> Context.generic val print_abs_rep: Proof.context -> unit - type quotients = {qtyp: typ, rtyp: typ, equiv_rel: term, equiv_thm: thm} + type quotients = {qtyp: typ, rtyp: typ, equiv_rel: term, equiv_thm: thm, quot_thm: thm} val transform_quotients: morphism -> quotients -> quotients val lookup_quotients: Proof.context -> string -> quotients option val lookup_quotients_global: theory -> string -> quotients option @@ -54,7 +54,7 @@ (* FIXME just one data slot (record) per program unit *) (* info about map- and rel-functions for a type *) -type quotmaps = {relmap: string} +type quotmaps = {relmap: string, quot_thm: thm} structure Quotmaps = Generic_Data ( @@ -71,19 +71,24 @@ val quotmaps_attribute_setup = Attrib.setup @{binding map} - ((Args.type_name true --| Scan.lift @{keyword "="}) -- Args.const_proper true >> - (fn (tyname, relname) => - let val minfo = {relmap = relname} + ((Args.type_name true --| Scan.lift @{keyword "="}) -- + (Scan.lift @{keyword "("} |-- Args.const_proper true --| Scan.lift @{keyword ","} -- + Attrib.thm --| Scan.lift @{keyword ")"}) >> + (fn (tyname, (relname, qthm)) => + let val minfo = {relmap = relname, quot_thm = qthm} in Thm.declaration_attribute (fn _ => Quotmaps.map (Symtab.update (tyname, minfo))) end)) "declaration of map information" fun print_quotmaps ctxt = let - fun prt_map (ty_name, {relmap}) = + fun prt_map (ty_name, {relmap, quot_thm}) = Pretty.block (separate (Pretty.brk 2) - (map Pretty.str - ["type:", ty_name, - "relation map:", relmap])) + [Pretty.str "type:", + Pretty.str ty_name, + Pretty.str "relation map:", + Pretty.str relmap, + Pretty.str "quot. theorem:", + Syntax.pretty_term ctxt (prop_of quot_thm)]) in map prt_map (Symtab.dest (Quotmaps.get (Context.Proof ctxt))) |> Pretty.big_list "maps for type constructors:" @@ -125,7 +130,7 @@ end (* info about quotient types *) -type quotients = {qtyp: typ, rtyp: typ, equiv_rel: term, equiv_thm: thm} +type quotients = {qtyp: typ, rtyp: typ, equiv_rel: term, equiv_thm: thm, quot_thm: thm} structure Quotients = Generic_Data ( @@ -135,11 +140,12 @@ fun merge data = Symtab.merge (K true) data ) -fun transform_quotients phi {qtyp, rtyp, equiv_rel, equiv_thm} = +fun transform_quotients phi {qtyp, rtyp, equiv_rel, equiv_thm, quot_thm} = {qtyp = Morphism.typ phi qtyp, rtyp = Morphism.typ phi rtyp, equiv_rel = Morphism.term phi equiv_rel, - equiv_thm = Morphism.thm phi equiv_thm} + equiv_thm = Morphism.thm phi equiv_thm, + quot_thm = Morphism.thm phi quot_thm} val lookup_quotients = Symtab.lookup o Quotients.get o Context.Proof val lookup_quotients_global = Symtab.lookup o Quotients.get o Context.Theory @@ -151,7 +157,7 @@ fun print_quotients ctxt = let - fun prt_quot {qtyp, rtyp, equiv_rel, equiv_thm} = + fun prt_quot {qtyp, rtyp, equiv_rel, equiv_thm, quot_thm} = Pretty.block (separate (Pretty.brk 2) [Pretty.str "quotient type:", Syntax.pretty_typ ctxt qtyp, @@ -160,7 +166,9 @@ Pretty.str "relation:", Syntax.pretty_term ctxt equiv_rel, Pretty.str "equiv. thm:", - Syntax.pretty_term ctxt (prop_of equiv_thm)]) + Syntax.pretty_term ctxt (prop_of equiv_thm), + Pretty.str "quot. thm:", + Syntax.pretty_term ctxt (prop_of quot_thm)]) in map (prt_quot o snd) (Symtab.dest (Quotients.get (Context.Proof ctxt))) |> Pretty.big_list "quotients:" diff -r 29e92b644d6c -r f067afe98049 src/HOL/Tools/Quotient/quotient_term.ML --- a/src/HOL/Tools/Quotient/quotient_term.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Tools/Quotient/quotient_term.ML Mon Mar 26 10:56:56 2012 +0200 @@ -20,6 +20,9 @@ val equiv_relation: Proof.context -> typ * typ -> term val equiv_relation_chk: Proof.context -> typ * typ -> term + val get_rel_from_quot_thm: thm -> term + val prove_quot_theorem: Proof.context -> typ * typ -> thm + val regularize_trm: Proof.context -> term * term -> term val regularize_trm_chk: Proof.context -> term * term -> term @@ -72,9 +75,6 @@ fun defined_mapfun_data ctxt s = Symtab.defined (Enriched_Type.entries ctxt) s - -(* makes a Free out of a TVar *) -fun mk_Free (TVar ((x, i), _)) = Free (unprefix "'" x ^ string_of_int i, dummyT) (* looks up the (varified) rty and qty for a quotient definition @@ -279,35 +279,10 @@ SOME map_data => Const (#relmap map_data, dummyT) | NONE => raise LIFT_MATCH ("get_relmap (no relation map function found for type " ^ s ^ ")")) -(* takes two type-environments and looks - up in both of them the variable v, which - must be listed in the environment -*) -fun double_lookup rtyenv qtyenv v = - let - val v' = fst (dest_TVar v) - in - (snd (the (Vartab.lookup rtyenv v')), snd (the (Vartab.lookup qtyenv v'))) - end - -fun mk_relmap ctxt vs rty = - let - val vs' = map (mk_Free) vs - - fun mk_relmap_aux rty = - case rty of - TVar _ => mk_Free rty - | Type (_, []) => HOLogic.eq_const rty - | Type (s, tys) => list_comb (get_relmap ctxt s, map mk_relmap_aux tys) - | _ => raise LIFT_MATCH ("mk_relmap (default)") - in - fold_rev Term.lambda vs' (mk_relmap_aux rty) - end - fun get_equiv_rel thy s = (case Quotient_Info.lookup_quotients thy s of SOME qdata => #equiv_rel qdata - | NONE => raise LIFT_MATCH ("get_quotdata (no quotient found for type " ^ s ^ ")")) + | NONE => raise LIFT_MATCH ("get_equiv_rel (no quotient found for type " ^ s ^ ")")) fun equiv_match_err ctxt ty_pat ty = let @@ -336,11 +311,10 @@ end else let - val (rty_pat, qty_pat as Type (_, vs)) = get_rty_qty ctxt s' - val rtyenv = match ctxt equiv_match_err rty_pat rty + val (Type (_, rtys), qty_pat) = get_rty_qty ctxt s' val qtyenv = match ctxt equiv_match_err qty_pat qty - val args_aux = map (double_lookup rtyenv qtyenv) vs - val args = map (equiv_relation ctxt) args_aux + val rtys' = map (Envir.subst_type qtyenv) rtys + val args = map (equiv_relation ctxt) (tys ~~ rtys') val eqv_rel = get_equiv_rel ctxt s' val eqv_rel' = force_typ ctxt eqv_rel ([rty, rty] ---> @{typ bool}) in @@ -348,8 +322,7 @@ then eqv_rel' else let - val rel_map = mk_relmap ctxt vs rty_pat - val result = list_comb (rel_map, args) + val result = list_comb (get_relmap ctxt s, args) in mk_rel_compose (result, eqv_rel') end @@ -361,6 +334,84 @@ equiv_relation ctxt (rty, qty) |> Syntax.check_term ctxt +(* generation of the Quotient theorem *) + +exception CODE_GEN of string + +fun get_quot_thm ctxt s = + let + val thy = Proof_Context.theory_of ctxt + in + (case Quotient_Info.lookup_quotients ctxt s of + SOME qdata => Thm.transfer thy (#quot_thm qdata) + | NONE => raise CODE_GEN ("No quotient type " ^ quote s ^ " found.")) + end + +fun get_rel_quot_thm ctxt s = + let + val thy = Proof_Context.theory_of ctxt + in + (case Quotient_Info.lookup_quotmaps ctxt s of + SOME map_data => Thm.transfer thy (#quot_thm map_data) + | NONE => raise CODE_GEN ("get_relmap (no relation map function found for type " ^ s ^ ")")) + end + +fun is_id_quot thm = (prop_of thm = prop_of @{thm identity_quotient}) + +infix 0 MRSL + +fun ants MRSL thm = fold (fn rl => fn thm => rl RS thm) ants thm + +exception NOT_IMPL of string + +fun get_rel_from_quot_thm quot_thm = + let + val (_ $ rel $ _ $ _) = (HOLogic.dest_Trueprop o prop_of) quot_thm + in + rel + end + +fun mk_quot_thm_compose (rel_quot_thm, quot_thm) = + let + val quot_thm_rel = get_rel_from_quot_thm quot_thm + in + if is_eq quot_thm_rel then [rel_quot_thm, quot_thm] MRSL @{thm OOO_eq_quotient} + else raise NOT_IMPL "nested quotients: not implemented yet" + end + +fun prove_quot_theorem ctxt (rty, qty) = + if rty = qty + then @{thm identity_quotient} + else + case (rty, qty) of + (Type (s, tys), Type (s', tys')) => + if s = s' + then + let + val args = map (prove_quot_theorem ctxt) (tys ~~ tys') + in + args MRSL (get_rel_quot_thm ctxt s) + end + else + let + val (Type (_, rtys), qty_pat) = get_rty_qty ctxt s' + val qtyenv = match ctxt equiv_match_err qty_pat qty + val rtys' = map (Envir.subst_type qtyenv) rtys + val args = map (prove_quot_theorem ctxt) (tys ~~ rtys') + val quot_thm = get_quot_thm ctxt s' + in + if forall is_id_quot args + then + quot_thm + else + let + val rel_quot_thm = args MRSL (get_rel_quot_thm ctxt s) + in + mk_quot_thm_compose (rel_quot_thm, quot_thm) + end + end + | _ => @{thm identity_quotient} + (*** Regularization ***) diff -r 29e92b644d6c -r f067afe98049 src/HOL/Tools/Quotient/quotient_type.ML --- a/src/HOL/Tools/Quotient/quotient_type.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Tools/Quotient/quotient_type.ML Mon Mar 26 10:56:56 2012 +0200 @@ -6,6 +6,8 @@ signature QUOTIENT_TYPE = sig + val can_generate_code_cert: thm -> bool + val add_quotient_type: ((string list * binding * mixfix) * (typ * term * bool) * ((binding * binding) option)) * thm -> local_theory -> Quotient_Info.quotients * local_theory @@ -25,7 +27,7 @@ val mem_def1 = @{lemma "y : Collect S ==> S y" by simp} val mem_def2 = @{lemma "S y ==> y : Collect S" by simp} -(* constructs the term lambda (c::rty => bool). EX (x::rty). c = rel x *) +(* constructs the term {c. EX (x::rty). rel x x \ c = Collect (rel x)} *) fun typedef_term rel rty lthy = let val [x, c] = @@ -76,6 +78,44 @@ Goal.prove lthy [] [] goal (K (typedef_quot_type_tac equiv_thm typedef_info)) end + +fun can_generate_code_cert quot_thm = + case Quotient_Term.get_rel_from_quot_thm quot_thm of + Const (@{const_name HOL.eq}, _) => true + | Const (@{const_name invariant}, _) $ _ => true + | _ => false + +fun define_abs_type quot_thm lthy = + if can_generate_code_cert quot_thm then + let + val abs_type_thm = quot_thm RS @{thm Quotient_abs_rep} + val add_abstype_attribute = + Thm.declaration_attribute (fn thm => Context.mapping (Code.add_abstype thm) I) + val add_abstype_attrib = Attrib.internal (K add_abstype_attribute); + in + lthy + |> (snd oo Local_Theory.note) ((Binding.empty, [add_abstype_attrib]), [abs_type_thm]) + end + else + lthy + +fun init_quotient_infr quot_thm equiv_thm lthy = + let + val (_ $ rel $ abs $ rep) = (HOLogic.dest_Trueprop o prop_of) quot_thm + val (qtyp, rtyp) = (dest_funT o fastype_of) rep + val qty_full_name = (fst o dest_Type) qtyp + val quotients = {qtyp = qtyp, rtyp = rtyp, equiv_rel = rel, equiv_thm = equiv_thm, + quot_thm = quot_thm } + fun quot_info phi = Quotient_Info.transform_quotients phi quotients + val abs_rep = {abs = abs, rep = rep} + fun abs_rep_info phi = Quotient_Info.transform_abs_rep phi abs_rep + in + lthy + |> Local_Theory.declaration {syntax = false, pervasive = true} + (fn phi => Quotient_Info.update_quotients qty_full_name (quot_info phi) + #> Quotient_Info.update_abs_rep qty_full_name (abs_rep_info phi)) + |> define_abs_type quot_thm + end (* main function for constructing a quotient type *) fun add_quotient_type (((vs, qty_name, mx), (rty, rel, partial), opt_morphs), equiv_thm) lthy = @@ -86,7 +126,7 @@ else equiv_thm RS @{thm equivp_implies_part_equivp} (* generates the typedef *) - val ((qty_full_name, typedef_info), lthy1) = + val ((_, typedef_info), lthy1) = typedef_make (vs, qty_name, mx, rel, rty) part_equiv lthy (* abs and rep functions from the typedef *) @@ -108,9 +148,9 @@ NONE => (Binding.prefix_name "rep_" qty_name, Binding.prefix_name "abs_" qty_name) | SOME morphs => morphs) - val ((abs_t, (_, abs_def)), lthy2) = lthy1 + val ((_, (_, abs_def)), lthy2) = lthy1 |> Local_Theory.define ((abs_name, NoSyn), ((Thm.def_binding abs_name, []), abs_trm)) - val ((rep_t, (_, rep_def)), lthy3) = lthy2 + val ((_, (_, rep_def)), lthy3) = lthy2 |> Local_Theory.define ((rep_name, NoSyn), ((Thm.def_binding rep_name, []), rep_trm)) (* quot_type theorem *) @@ -126,15 +166,11 @@ val equiv_thm_name = Binding.suffix_name "_equivp" qty_name (* storing the quotients *) - val quotients = {qtyp = Abs_ty, rtyp = rty, equiv_rel = rel, equiv_thm = equiv_thm} - - fun qinfo phi = Quotient_Info.transform_quotients phi quotients - fun abs_rep phi = Quotient_Info.transform_abs_rep phi {abs = abs_t, rep = rep_t} + val quotients = {qtyp = Abs_ty, rtyp = rty, equiv_rel = rel, equiv_thm = equiv_thm, + quot_thm = quotient_thm} val lthy4 = lthy3 - |> Local_Theory.declaration {syntax = false, pervasive = true} - (fn phi => Quotient_Info.update_quotients qty_full_name (qinfo phi) - #> Quotient_Info.update_abs_rep qty_full_name (abs_rep phi)) + |> init_quotient_infr quotient_thm equiv_thm |> (snd oo Local_Theory.note) ((equiv_thm_name, if partial then [] else [Attrib.internal (K Quotient_Info.equiv_rules_add)]), @@ -264,15 +300,43 @@ val partial = Scan.optional (Parse.reserved "partial" -- @{keyword ":"} >> K true) false +val quotspec_parser = + Parse.and_list1 + ((Parse.type_args -- Parse.binding) -- + (* FIXME Parse.type_args_constrained and standard treatment of sort constraints *) + Parse.opt_mixfix -- (@{keyword "="} |-- Parse.typ) -- + (@{keyword "/"} |-- (partial -- Parse.term)) -- + Scan.option (@{keyword "morphisms"} |-- Parse.!!! (Parse.binding -- Parse.binding))) + val _ = Outer_Syntax.local_theory_to_proof @{command_spec "quotient_type"} "quotient type definitions (require equivalence proofs)" - (Parse.and_list1 - ((Parse.type_args -- Parse.binding) -- - (* FIXME Parse.type_args_constrained and standard treatment of sort constraints *) - Parse.opt_mixfix -- (@{keyword "="} |-- Parse.typ) -- - (@{keyword "/"} |-- (partial -- Parse.term)) -- - Scan.option (@{keyword "morphisms"} |-- Parse.!!! (Parse.binding -- Parse.binding))) - >> quotient_type_cmd) + (quotspec_parser >> quotient_type_cmd) + +(* Setup lifting using type_def_thm *) + +exception SETUP_LIFT_TYPE of string + +fun setup_lift_type typedef_thm = + let + val typedef_set = (snd o dest_comb o HOLogic.dest_Trueprop o prop_of) typedef_thm + val (quot_thm, equivp_thm) = (case typedef_set of + Const ("Orderings.top_class.top", _) => + (typedef_thm RS @{thm copy_type_to_Quotient}, + typedef_thm RS @{thm copy_type_to_equivp}) + | Const (@{const_name "Collect"}, _) $ Abs (_, _, _ $ Bound 0) => + (typedef_thm RS @{thm invariant_type_to_Quotient}, + typedef_thm RS @{thm invariant_type_to_part_equivp}) + | _ => raise SETUP_LIFT_TYPE "unsupported typedef theorem") + in + init_quotient_infr quot_thm equivp_thm + end + +fun setup_lift_type_aux xthm lthy = setup_lift_type (singleton (Attrib.eval_thms lthy) xthm) lthy + +val _ = + Outer_Syntax.local_theory @{command_spec "setup_lifting"} + "Setup lifting infrastracture" + (Parse_Spec.xthm >> (fn xthm => setup_lift_type_aux xthm)) end; diff -r 29e92b644d6c -r f067afe98049 src/HOL/Tools/SMT/smt_normalize.ML --- a/src/HOL/Tools/SMT/smt_normalize.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Tools/SMT/smt_normalize.ML Mon Mar 26 10:56:56 2012 +0200 @@ -451,7 +451,7 @@ val nat_ops = simple_nat_ops @ mult_nat_ops - val nat_consts = nat_ops @ [@{const number_of (nat)}, + val nat_consts = nat_ops @ [@{const numeral (nat)}, @{const zero_class.zero (nat)}, @{const one_class.one (nat)}] val nat_int_coercions = [@{const of_nat (int)}, @{const nat}] @@ -466,7 +466,7 @@ val expands = map mk_meta_eq @{lemma "0 = nat 0" "1 = nat 1" - "(number_of :: int => nat) = (%i. nat (number_of i))" + "(numeral :: num => nat) = (%i. nat (numeral i))" "op < = (%a b. int a < int b)" "op <= = (%a b. int a <= int b)" "Suc = (%a. nat (int a + 1))" @@ -493,8 +493,7 @@ let val eq = SMT_Utils.mk_cequals lhs (Numeral.mk_cnumber @{ctyp int} i) val ss = HOL_ss - addsimps [@{thm Nat_Numeral.int_nat_number_of}] - addsimps @{thms neg_simps} + addsimps [@{thm Nat_Numeral.int_numeral}] fun tac _ = Simplifier.simp_tac (Simplifier.context ctxt ss) 1 in Goal.norm_result (Goal.prove_internal [] eq tac) end @@ -507,7 +506,7 @@ fun int_conv ctxt ct = (case Thm.term_of ct of - @{const of_nat (int)} $ (n as (@{const number_of (nat)} $ _)) => + @{const of_nat (int)} $ (n as (@{const numeral (nat)} $ _)) => Conv.rewr_conv (mk_number_eq ctxt (snd (HOLogic.dest_number n)) ct) | @{const of_nat (int)} $ _ => (Conv.rewrs_conv ints then_conv Conv.sub_conv ints_conv ctxt) else_conv @@ -549,23 +548,15 @@ rewrite Numeral1 into 1 *) - fun is_strange_number ctxt (t as Const (@{const_name number_of}, _) $ _) = + fun is_strange_number ctxt (t as Const (@{const_name neg_numeral}, _) $ _) = (case try HOLogic.dest_number t of SOME (_, i) => SMT_Builtin.is_builtin_num ctxt t andalso i < 2 | NONE => false) | is_strange_number _ _ = false val pos_num_ss = HOL_ss - addsimps [@{thm Int.number_of_minus}, @{thm Int.number_of_Min}] - addsimps [@{thm Int.number_of_Pls}, @{thm Int.numeral_1_eq_1}] - addsimps @{thms Int.pred_bin_simps} - addsimps @{thms Int.normalize_bin_simps} - addsimps @{lemma - "Int.Min = - Int.Bit1 Int.Pls" - "Int.Bit0 (- Int.Pls) = - Int.Pls" - "Int.Bit0 (- k) = - Int.Bit0 k" - "Int.Bit1 (- k) = - Int.Bit1 (Int.pred k)" - by simp_all (simp add: pred_def)} + addsimps [@{thm Num.numeral_One}] + addsimps [@{thm Num.neg_numeral_def}] fun norm_num_conv ctxt = SMT_Utils.if_conv (is_strange_number ctxt) diff -r 29e92b644d6c -r f067afe98049 src/HOL/Tools/SMT/z3_proof_tools.ML --- a/src/HOL/Tools/SMT/z3_proof_tools.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Tools/SMT/z3_proof_tools.ML Mon Mar 26 10:56:56 2012 +0200 @@ -334,14 +334,12 @@ val basic_simpset = HOL_ss addsimps @{thms field_simps} addsimps [@{thm times_divide_eq_right}, @{thm times_divide_eq_left}] - addsimps @{thms arith_special} addsimps @{thms less_bin_simps} - addsimps @{thms le_bin_simps} addsimps @{thms eq_bin_simps} - addsimps @{thms add_bin_simps} addsimps @{thms succ_bin_simps} - addsimps @{thms minus_bin_simps} addsimps @{thms pred_bin_simps} - addsimps @{thms mult_bin_simps} addsimps @{thms iszero_simps} + addsimps @{thms arith_special} addsimps @{thms arith_simps} + addsimps @{thms rel_simps} addsimps @{thms array_rules} addsimps @{thms term_true_def} addsimps @{thms term_false_def} addsimps @{thms z3div_def} addsimps @{thms z3mod_def} + addsimprocs [@{simproc binary_int_div}, @{simproc binary_int_mod}] addsimprocs [ Simplifier.simproc_global @{theory} "fast_int_arith" [ "(m::int) < n", "(m::int) <= n", "(m::int) = n"] (K Lin_Arith.simproc), diff -r 29e92b644d6c -r f067afe98049 src/HOL/Tools/arith_data.ML --- a/src/HOL/Tools/arith_data.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Tools/arith_data.ML Mon Mar 26 10:56:56 2012 +0200 @@ -68,7 +68,8 @@ (* some specialized syntactic operations *) -fun mk_number T n = HOLogic.number_of_const T $ HOLogic.mk_numeral n; +fun mk_number T 1 = HOLogic.numeral_const T $ HOLogic.one_const + | mk_number T n = HOLogic.mk_number T n; val mk_plus = HOLogic.mk_binop @{const_name Groups.plus}; diff -r 29e92b644d6c -r f067afe98049 src/HOL/Tools/float_syntax.ML --- a/src/HOL/Tools/float_syntax.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Tools/float_syntax.ML Mon Mar 26 10:56:56 2012 +0200 @@ -18,12 +18,15 @@ fun mk_number i = let - fun mk 0 = Syntax.const @{const_syntax Int.Pls} - | mk ~1 = Syntax.const @{const_syntax Int.Min} + fun mk 1 = Syntax.const @{const_syntax Num.One} | mk i = let val (q, r) = Integer.div_mod i 2 in HOLogic.mk_bit r $ (mk q) end; - in Syntax.const @{const_syntax Int.number_of} $ mk i end; + in + if i = 0 then Syntax.const @{const_syntax Groups.zero} + else if i > 0 then Syntax.const @{const_syntax Num.numeral} $ mk i + else Syntax.const @{const_syntax Num.neg_numeral} $ mk (~i) + end; fun mk_frac str = let diff -r 29e92b644d6c -r f067afe98049 src/HOL/Tools/hologic.ML --- a/src/HOL/Tools/hologic.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Tools/hologic.ML Mon Mar 26 10:56:56 2012 +0200 @@ -93,15 +93,15 @@ val size_const: typ -> term val code_numeralT: typ val intT: typ - val pls_const: term - val min_const: term + val one_const: term val bit0_const: term val bit1_const: term val mk_bit: int -> term val dest_bit: term -> int val mk_numeral: int -> term - val dest_numeral: term -> int - val number_of_const: typ -> term + val dest_num: term -> int + val numeral_const: typ -> term + val neg_numeral_const: typ -> term val add_numerals: term -> (term * typ) list -> (term * typ) list val mk_number: typ -> int -> term val dest_number: term -> typ * int @@ -492,50 +492,54 @@ val code_numeralT = Type ("Code_Numeral.code_numeral", []); -(* binary numerals and int -- non-unique representation due to leading zeros/ones! *) +(* binary numerals and int *) +val numT = Type ("Num.num", []); val intT = Type ("Int.int", []); -val pls_const = Const ("Int.Pls", intT) -and min_const = Const ("Int.Min", intT) -and bit0_const = Const ("Int.Bit0", intT --> intT) -and bit1_const = Const ("Int.Bit1", intT --> intT); +val one_const = Const ("Num.num.One", numT) +and bit0_const = Const ("Num.num.Bit0", numT --> numT) +and bit1_const = Const ("Num.num.Bit1", numT --> numT); fun mk_bit 0 = bit0_const | mk_bit 1 = bit1_const | mk_bit _ = raise TERM ("mk_bit", []); -fun dest_bit (Const ("Int.Bit0", _)) = 0 - | dest_bit (Const ("Int.Bit1", _)) = 1 +fun dest_bit (Const ("Num.num.Bit0", _)) = 0 + | dest_bit (Const ("Num.num.Bit1", _)) = 1 | dest_bit t = raise TERM ("dest_bit", [t]); -fun mk_numeral 0 = pls_const - | mk_numeral ~1 = min_const - | mk_numeral i = - let val (q, r) = Integer.div_mod i 2; - in mk_bit r $ mk_numeral q end; +fun mk_numeral i = + let fun mk 1 = one_const + | mk i = let val (q, r) = Integer.div_mod i 2 in mk_bit r $ mk q end + in if i > 0 then mk i else raise TERM ("mk_numeral: " ^ string_of_int i, []) + end -fun dest_numeral (Const ("Int.Pls", _)) = 0 - | dest_numeral (Const ("Int.Min", _)) = ~1 - | dest_numeral (Const ("Int.Bit0", _) $ bs) = 2 * dest_numeral bs - | dest_numeral (Const ("Int.Bit1", _) $ bs) = 2 * dest_numeral bs + 1 - | dest_numeral t = raise TERM ("dest_numeral", [t]); +fun dest_num (Const ("Num.num.One", _)) = 1 + | dest_num (Const ("Num.num.Bit0", _) $ bs) = 2 * dest_num bs + | dest_num (Const ("Num.num.Bit1", _) $ bs) = 2 * dest_num bs + 1 + | dest_num t = raise TERM ("dest_num", [t]); -fun number_of_const T = Const ("Int.number_class.number_of", intT --> T); +fun numeral_const T = Const ("Num.numeral_class.numeral", numT --> T); +fun neg_numeral_const T = Const ("Num.neg_numeral_class.neg_numeral", numT --> T); -fun add_numerals (Const ("Int.number_class.number_of", Type (_, [_, T])) $ t) = cons (t, T) +fun add_numerals (Const ("Num.numeral_class.numeral", Type (_, [_, T])) $ t) = cons (t, T) | add_numerals (t $ u) = add_numerals t #> add_numerals u | add_numerals (Abs (_, _, t)) = add_numerals t | add_numerals _ = I; fun mk_number T 0 = Const ("Groups.zero_class.zero", T) | mk_number T 1 = Const ("Groups.one_class.one", T) - | mk_number T i = number_of_const T $ mk_numeral i; + | mk_number T i = + if i > 0 then numeral_const T $ mk_numeral i + else neg_numeral_const T $ mk_numeral (~ i); fun dest_number (Const ("Groups.zero_class.zero", T)) = (T, 0) | dest_number (Const ("Groups.one_class.one", T)) = (T, 1) - | dest_number (Const ("Int.number_class.number_of", Type ("fun", [_, T])) $ t) = - (T, dest_numeral t) + | dest_number (Const ("Num.numeral_class.numeral", Type ("fun", [_, T])) $ t) = + (T, dest_num t) + | dest_number (Const ("Num.neg_numeral_class.neg_numeral", Type ("fun", [_, T])) $ t) = + (T, ~ (dest_num t)) | dest_number t = raise TERM ("dest_number", [t]); diff -r 29e92b644d6c -r f067afe98049 src/HOL/Tools/int_arith.ML --- a/src/HOL/Tools/int_arith.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Tools/int_arith.ML Mon Mar 26 10:56:56 2012 +0200 @@ -78,7 +78,7 @@ proc = sproc, identifier = []} fun number_of thy T n = - if not (Sign.of_sort thy (T, @{sort number})) + if not (Sign.of_sort thy (T, @{sort numeral})) then raise CTERM ("number_of", []) else Numeral.mk_cnumber (Thm.ctyp_of thy T) n; diff -r 29e92b644d6c -r f067afe98049 src/HOL/Tools/lin_arith.ML --- a/src/HOL/Tools/lin_arith.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Tools/lin_arith.ML Mon Mar 26 10:56:56 2012 +0200 @@ -174,14 +174,17 @@ | (NONE, m') => apsnd (Rat.mult (Rat.inv m')) (demult (s, m))) (* terms that evaluate to numeric constants *) | demult (Const (@{const_name Groups.uminus}, _) $ t, m) = demult (t, Rat.neg m) - | demult (Const (@{const_name Groups.zero}, _), m) = (NONE, Rat.zero) + | demult (Const (@{const_name Groups.zero}, _), _) = (NONE, Rat.zero) | demult (Const (@{const_name Groups.one}, _), m) = (NONE, m) - (*Warning: in rare cases number_of encloses a non-numeral, - in which case dest_numeral raises TERM; hence all the handles below. + (*Warning: in rare cases (neg_)numeral encloses a non-numeral, + in which case dest_num raises TERM; hence all the handles below. Same for Suc-terms that turn out not to be numerals - although the simplifier should eliminate those anyway ...*) - | demult (t as Const ("Int.number_class.number_of", _) $ n, m) = - ((NONE, Rat.mult m (Rat.rat_of_int (HOLogic.dest_numeral n))) + | demult (t as Const ("Num.numeral_class.numeral", _) $ n, m) = + ((NONE, Rat.mult m (Rat.rat_of_int (HOLogic.dest_num n))) + handle TERM _ => (SOME t, m)) + | demult (t as Const ("Num.neg_numeral_class.neg_numeral", _) $ n, m) = + ((NONE, Rat.mult m (Rat.rat_of_int (~ (HOLogic.dest_num n)))) handle TERM _ => (SOME t, m)) | demult (t as Const (@{const_name Suc}, _) $ _, m) = ((NONE, Rat.mult m (Rat.rat_of_int (HOLogic.dest_nat t))) @@ -219,10 +222,13 @@ (case demult inj_consts (all, m) of (NONE, m') => (p, Rat.add i m') | (SOME u, m') => add_atom u m' pi) - | poly (all as Const ("Int.number_class.number_of", Type(_,[_,T])) $ t, m, pi as (p, i)) = - (let val k = HOLogic.dest_numeral t - val k2 = if k < 0 andalso T = HOLogic.natT then 0 else k - in (p, Rat.add i (Rat.mult m (Rat.rat_of_int k2))) end + | poly (all as Const ("Num.numeral_class.numeral", Type(_,[_,_])) $ t, m, pi as (p, i)) = + (let val k = HOLogic.dest_num t + in (p, Rat.add i (Rat.mult m (Rat.rat_of_int k))) end + handle TERM _ => add_atom all m pi) + | poly (all as Const ("Num.neg_numeral_class.neg_numeral", Type(_,[_,_])) $ t, m, pi as (p, i)) = + (let val k = HOLogic.dest_num t + in (p, Rat.add i (Rat.mult m (Rat.rat_of_int (~ k)))) end handle TERM _ => add_atom all m pi) | poly (all as Const f $ x, m, pi) = if member (op =) inj_consts f then poly (x, m, pi) else add_atom all m pi @@ -464,9 +470,9 @@ in SOME [(HOLogic.natT :: Ts, subgoal1), (Ts, subgoal2)] end - (* ?P ((?n::nat) mod (number_of ?k)) = - ((number_of ?k = 0 --> ?P ?n) & (~ (number_of ?k = 0) --> - (ALL i j. j < number_of ?k --> ?n = number_of ?k * i + j --> ?P j))) *) + (* ?P ((?n::nat) mod (numeral ?k)) = + ((numeral ?k = 0 --> ?P ?n) & (~ (numeral ?k = 0) --> + (ALL i j. j < numeral ?k --> ?n = numeral ?k * i + j --> ?P j))) *) | (Const ("Divides.div_class.mod", Type ("fun", [@{typ nat}, _])), [t1, t2]) => let val rev_terms = rev terms @@ -496,9 +502,9 @@ in SOME [(Ts, subgoal1), (split_type :: split_type :: Ts, subgoal2)] end - (* ?P ((?n::nat) div (number_of ?k)) = - ((number_of ?k = 0 --> ?P 0) & (~ (number_of ?k = 0) --> - (ALL i j. j < number_of ?k --> ?n = number_of ?k * i + j --> ?P i))) *) + (* ?P ((?n::nat) div (numeral ?k)) = + ((numeral ?k = 0 --> ?P 0) & (~ (numeral ?k = 0) --> + (ALL i j. j < numeral ?k --> ?n = numeral ?k * i + j --> ?P i))) *) | (Const ("Divides.div_class.div", Type ("fun", [@{typ nat}, _])), [t1, t2]) => let val rev_terms = rev terms @@ -528,14 +534,14 @@ in SOME [(Ts, subgoal1), (split_type :: split_type :: Ts, subgoal2)] end - (* ?P ((?n::int) mod (number_of ?k)) = - ((number_of ?k = 0 --> ?P ?n) & - (0 < number_of ?k --> + (* ?P ((?n::int) mod (numeral ?k)) = + ((numeral ?k = 0 --> ?P ?n) & + (0 < numeral ?k --> (ALL i j. - 0 <= j & j < number_of ?k & ?n = number_of ?k * i + j --> ?P j)) & - (number_of ?k < 0 --> + 0 <= j & j < numeral ?k & ?n = numeral ?k * i + j --> ?P j)) & + (numeral ?k < 0 --> (ALL i j. - number_of ?k < j & j <= 0 & ?n = number_of ?k * i + j --> ?P j))) *) + numeral ?k < j & j <= 0 & ?n = numeral ?k * i + j --> ?P j))) *) | (Const ("Divides.div_class.mod", Type ("fun", [Type ("Int.int", []), _])), [t1, t2]) => let @@ -582,14 +588,14 @@ in SOME [(Ts, subgoal1), (Ts', subgoal2), (Ts', subgoal3)] end - (* ?P ((?n::int) div (number_of ?k)) = - ((number_of ?k = 0 --> ?P 0) & - (0 < number_of ?k --> + (* ?P ((?n::int) div (numeral ?k)) = + ((numeral ?k = 0 --> ?P 0) & + (0 < numeral ?k --> (ALL i j. - 0 <= j & j < number_of ?k & ?n = number_of ?k * i + j --> ?P i)) & - (number_of ?k < 0 --> + 0 <= j & j < numeral ?k & ?n = numeral ?k * i + j --> ?P i)) & + (numeral ?k < 0 --> (ALL i j. - number_of ?k < j & j <= 0 & ?n = number_of ?k * i + j --> ?P i))) *) + numeral ?k < j & j <= 0 & ?n = numeral ?k * i + j --> ?P i))) *) | (Const ("Divides.div_class.div", Type ("fun", [Type ("Int.int", []), _])), [t1, t2]) => let diff -r 29e92b644d6c -r f067afe98049 src/HOL/Tools/nat_numeral_simprocs.ML --- a/src/HOL/Tools/nat_numeral_simprocs.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Tools/nat_numeral_simprocs.ML Mon Mar 26 10:56:56 2012 +0200 @@ -25,15 +25,16 @@ structure Nat_Numeral_Simprocs : NAT_NUMERAL_SIMPROCS = struct -(*Maps n to #n for n = 0, 1, 2*) -val numeral_syms = [@{thm nat_numeral_0_eq_0} RS sym, @{thm nat_numeral_1_eq_1} RS sym, @{thm numeral_2_eq_2} RS sym]; +(*Maps n to #n for n = 1, 2*) +val numeral_syms = [@{thm numeral_1_eq_1} RS sym, @{thm numeral_2_eq_2} RS sym]; val numeral_sym_ss = HOL_basic_ss addsimps numeral_syms; val rename_numerals = simplify numeral_sym_ss o Thm.transfer @{theory}; (*Utilities*) -fun mk_number n = HOLogic.number_of_const HOLogic.natT $ HOLogic.mk_numeral n; +fun mk_number 1 = HOLogic.numeral_const HOLogic.natT $ HOLogic.one_const + | mk_number n = HOLogic.mk_number HOLogic.natT n; fun dest_number t = Int.max (0, snd (HOLogic.dest_number t)); fun find_first_numeral past (t::terms) = @@ -59,14 +60,13 @@ (** Other simproc items **) val bin_simps = - [@{thm nat_numeral_0_eq_0} RS sym, @{thm nat_numeral_1_eq_1} RS sym, - @{thm add_nat_number_of}, @{thm nat_number_of_add_left}, - @{thm diff_nat_number_of}, @{thm le_number_of_eq_not_less}, - @{thm mult_nat_number_of}, @{thm nat_number_of_mult_left}, - @{thm less_nat_number_of}, + [@{thm numeral_1_eq_1} RS sym, + @{thm numeral_plus_numeral}, @{thm add_numeral_left}, + @{thm diff_nat_numeral}, @{thm diff_0_eq_0}, @{thm diff_0}, + @{thm numeral_times_numeral}, @{thm mult_numeral_left(1)}, @{thm if_True}, @{thm if_False}, @{thm not_False_eq_True}, - @{thm Let_number_of}, @{thm nat_number_of}] @ - @{thms arith_simps} @ @{thms rel_simps} @ @{thms neg_simps}; + @{thm nat_0}, @{thm nat_numeral}, @{thm nat_neg_numeral}] @ + @{thms arith_simps} @ @{thms rel_simps}; (*** CancelNumerals simprocs ***) @@ -115,7 +115,7 @@ handle TERM _ => (k, t::ts); (*Code for testing whether numerals are already used in the goal*) -fun is_numeral (Const(@{const_name Int.number_of}, _) $ w) = true +fun is_numeral (Const(@{const_name Num.numeral}, _) $ w) = true | is_numeral _ = false; fun prod_has_numeral t = exists is_numeral (dest_prod t); @@ -147,7 +147,7 @@ val simplify_meta_eq = Arith_Data.simplify_meta_eq - ([@{thm nat_numeral_0_eq_0}, @{thm numeral_1_eq_Suc_0}, @{thm Nat.add_0}, @{thm Nat.add_0_right}, + ([@{thm numeral_1_eq_Suc_0}, @{thm Nat.add_0}, @{thm Nat.add_0_right}, @{thm mult_0}, @{thm mult_0_right}, @{thm mult_1}, @{thm mult_1_right}] @ contra_rules); diff -r 29e92b644d6c -r f067afe98049 src/HOL/Tools/numeral.ML --- a/src/HOL/Tools/numeral.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Tools/numeral.ML Mon Mar 26 10:56:56 2012 +0200 @@ -16,16 +16,20 @@ (* numeral *) -fun mk_cbit 0 = @{cterm "Int.Bit0"} - | mk_cbit 1 = @{cterm "Int.Bit1"} +fun mk_cbit 0 = @{cterm "Num.Bit0"} + | mk_cbit 1 = @{cterm "Num.Bit1"} | mk_cbit _ = raise CTERM ("mk_cbit", []); -fun mk_cnumeral 0 = @{cterm "Int.Pls"} - | mk_cnumeral ~1 = @{cterm "Int.Min"} - | mk_cnumeral i = +fun mk_cnumeral i = + let + fun mk 1 = @{cterm "Num.One"} + | mk i = let val (q, r) = Integer.div_mod i 2 in - Thm.apply (mk_cbit r) (mk_cnumeral q) - end; + Thm.apply (mk_cbit r) (mk q) + end + in + if i > 0 then mk i else raise CTERM ("mk_cnumeral: negative input", []) + end (* number *) @@ -38,8 +42,11 @@ val one = @{cpat "1"}; val oneT = Thm.ctyp_of_term one; -val number_of = @{cpat "number_of"}; -val numberT = Thm.ctyp_of @{theory} (Term.range_type (Thm.typ_of (Thm.ctyp_of_term number_of))); +val numeral = @{cpat "numeral"}; +val numeralT = Thm.ctyp_of @{theory} (Term.range_type (Thm.typ_of (Thm.ctyp_of_term numeral))); + +val neg_numeral = @{cpat "neg_numeral"}; +val neg_numeralT = Thm.ctyp_of @{theory} (Term.range_type (Thm.typ_of (Thm.ctyp_of_term neg_numeral))); fun instT T V = Thm.instantiate_cterm ([(V, T)], []); @@ -47,7 +54,9 @@ fun mk_cnumber T 0 = instT T zeroT zero | mk_cnumber T 1 = instT T oneT one - | mk_cnumber T i = Thm.apply (instT T numberT number_of) (mk_cnumeral i); + | mk_cnumber T i = + if i > 0 then Thm.apply (instT T numeralT numeral) (mk_cnumeral i) + else Thm.apply (instT T neg_numeralT neg_numeral) (mk_cnumeral (~i)); end; @@ -58,27 +67,23 @@ fun add_code number_of negative print target thy = let - fun dest_numeral pls' min' bit0' bit1' thm = + fun dest_numeral one' bit0' bit1' thm t = let fun dest_bit (IConst (c, _)) = if c = bit0' then 0 else if c = bit1' then 1 else Code_Printer.eqn_error thm "Illegal numeral expression: illegal bit" | dest_bit _ = Code_Printer.eqn_error thm "Illegal numeral expression: illegal bit"; - fun dest_num (IConst (c, _)) = if c = pls' then SOME 0 - else if c = min' then - if negative then SOME ~1 else NONE + fun dest_num (IConst (c, _)) = if c = one' then 1 else Code_Printer.eqn_error thm "Illegal numeral expression: illegal leading digit" - | dest_num (t1 `$ t2) = - let val (n, b) = (dest_num t2, dest_bit t1) - in case n of SOME n => SOME (2 * n + b) | NONE => NONE end + | dest_num (t1 `$ t2) = 2 * dest_num t2 + dest_bit t1 | dest_num _ = Code_Printer.eqn_error thm "Illegal numeral expression: illegal term"; - in dest_num end; - fun pretty literals [pls', min', bit0', bit1'] _ thm _ _ [(t, _)] = - (Code_Printer.str o print literals o the_default 0 o dest_numeral pls' min' bit0' bit1' thm) t; + in if negative then ~ (dest_num t) else dest_num t end; + fun pretty literals [one', bit0', bit1'] _ thm _ _ [(t, _)] = + (Code_Printer.str o print literals o dest_numeral one' bit0' bit1' thm) t; in thy |> Code_Target.add_const_syntax target number_of - (SOME (Code_Printer.complex_const_syntax (1, ([@{const_name Int.Pls}, @{const_name Int.Min}, - @{const_name Int.Bit0}, @{const_name Int.Bit1}], pretty)))) + (SOME (Code_Printer.complex_const_syntax (1, ([@{const_name Num.One}, + @{const_name Num.Bit0}, @{const_name Num.Bit1}], pretty)))) end; end; (*local*) diff -r 29e92b644d6c -r f067afe98049 src/HOL/Tools/numeral_simprocs.ML --- a/src/HOL/Tools/numeral_simprocs.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Tools/numeral_simprocs.ML Mon Mar 26 10:56:56 2012 +0200 @@ -66,6 +66,7 @@ (* build product with trailing 1 rather than Numeral 1 in order to avoid the unnecessary restriction to type class number_ring which is not required for cancellation of common factors in divisions. + UPDATE: this reasoning no longer applies (number_ring is gone) *) fun mk_prod T = let val one = one_of T @@ -148,22 +149,24 @@ (*This resembles Term_Ord.term_ord, but it puts binary numerals before other non-atomic terms.*) -local open Term -in -fun numterm_ord (Abs (_, T, t), Abs(_, U, u)) = - (case numterm_ord (t, u) of EQUAL => Term_Ord.typ_ord (T, U) | ord => ord) - | numterm_ord - (Const(@{const_name Int.number_of}, _) $ v, Const(@{const_name Int.number_of}, _) $ w) = - num_ord (HOLogic.dest_numeral v, HOLogic.dest_numeral w) - | numterm_ord (Const(@{const_name Int.number_of}, _) $ _, _) = LESS - | numterm_ord (_, Const(@{const_name Int.number_of}, _) $ _) = GREATER - | numterm_ord (t, u) = - (case int_ord (size_of_term t, size_of_term u) of - EQUAL => +local open Term +in +fun numterm_ord (t, u) = + case (try HOLogic.dest_number t, try HOLogic.dest_number u) of + (SOME (_, i), SOME (_, j)) => num_ord (i, j) + | (SOME _, NONE) => LESS + | (NONE, SOME _) => GREATER + | _ => ( + case (t, u) of + (Abs (_, T, t), Abs(_, U, u)) => + (prod_ord numterm_ord Term_Ord.typ_ord ((t, T), (u, U))) + | _ => ( + case int_ord (size_of_term t, size_of_term u) of + EQUAL => let val (f, ts) = strip_comb t and (g, us) = strip_comb u in - (case Term_Ord.hd_ord (f, g) of EQUAL => numterms_ord (ts, us) | ord => ord) + (prod_ord Term_Ord.hd_ord numterms_ord ((f, ts), (g, us))) end - | ord => ord) + | ord => ord)) and numterms_ord (ts, us) = list_ord numterm_ord (ts, us) end; @@ -171,16 +174,16 @@ val num_ss = HOL_basic_ss |> Simplifier.set_termless numtermless; -(*Maps 0 to Numeral0 and 1 to Numeral1 so that arithmetic isn't complicated by the abstract 0 and 1.*) -val numeral_syms = [@{thm numeral_0_eq_0} RS sym, @{thm numeral_1_eq_1} RS sym]; +(*Maps 1 to Numeral1 so that arithmetic isn't complicated by the abstract 1.*) +val numeral_syms = [@{thm numeral_1_eq_1} RS sym]; -(*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1, 1*x, x*1, x/1 *) +(*Simplify 0+n, n+0, Numeral1*n, n*Numeral1, 1*x, x*1, x/1 *) val add_0s = @{thms add_0s}; val mult_1s = @{thms mult_1s mult_1_left mult_1_right divide_1}; (* For post-simplification of the rhs of simproc-generated rules *) val post_simps = - [@{thm numeral_0_eq_0}, @{thm numeral_1_eq_1}, + [@{thm numeral_1_eq_1}, @{thm add_0_left}, @{thm add_0_right}, @{thm mult_zero_left}, @{thm mult_zero_right}, @{thm mult_1_left}, @{thm mult_1_right}, @@ -195,18 +198,24 @@ (*To perform binary arithmetic. The "left" rewriting handles patterns created by the Numeral_Simprocs, such as 3 * (5 * x). *) -val simps = [@{thm numeral_0_eq_0} RS sym, @{thm numeral_1_eq_1} RS sym, - @{thm add_number_of_left}, @{thm mult_number_of_left}] @ - @{thms arith_simps} @ @{thms rel_simps}; - +val simps = + [@{thm numeral_1_eq_1} RS sym] @ + @{thms add_numeral_left} @ + @{thms add_neg_numeral_left} @ + @{thms mult_numeral_left} @ + @{thms arith_simps} @ @{thms rel_simps}; + (*Binary arithmetic BUT NOT ADDITION since it may collapse adjacent terms during re-arrangement*) val non_add_simps = - subtract Thm.eq_thm [@{thm add_number_of_left}, @{thm number_of_add} RS sym] simps; + subtract Thm.eq_thm + (@{thms add_numeral_left} @ + @{thms add_neg_numeral_left} @ + @{thms numeral_plus_numeral} @ + @{thms add_neg_numeral_simps}) simps; (*To evaluate binary negations of coefficients*) -val minus_simps = [@{thm numeral_m1_eq_minus_1} RS sym, @{thm number_of_minus} RS sym] @ - @{thms minus_bin_simps} @ @{thms pred_bin_simps}; +val minus_simps = [@{thm minus_zero}, @{thm minus_one}, @{thm minus_numeral}, @{thm minus_neg_numeral}]; (*To let us treat subtraction as addition*) val diff_simps = [@{thm diff_minus}, @{thm minus_add_distrib}, @{thm minus_minus}]; @@ -365,9 +374,7 @@ (* simp_thms are necessary because some of the cancellation rules below (e.g. mult_less_cancel_left) introduce various logical connectives *) - val numeral_simp_ss = HOL_basic_ss addsimps - [@{thm eq_number_of_eq}, @{thm less_number_of}, @{thm le_number_of}] @ simps - @ @{thms simp_thms} + val numeral_simp_ss = HOL_basic_ss addsimps simps @ @{thms simp_thms} fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) val simplify_meta_eq = Arith_Data.simplify_meta_eq ([@{thm Nat.add_0}, @{thm Nat.add_0_right}] @ post_simps) @@ -425,13 +432,16 @@ val field_cancel_numeral_factors = map (prep_simproc @{theory}) [("field_eq_cancel_numeral_factor", - ["(l::'a::{field,number_ring}) * m = n", - "(l::'a::{field,number_ring}) = m * n"], + ["(l::'a::field) * m = n", + "(l::'a::field) = m * n"], K EqCancelNumeralFactor.proc), ("field_cancel_numeral_factor", - ["((l::'a::{field_inverse_zero,number_ring}) * m) / n", - "(l::'a::{field_inverse_zero,number_ring}) / (m * n)", - "((number_of v)::'a::{field_inverse_zero,number_ring}) / (number_of w)"], + ["((l::'a::field_inverse_zero) * m) / n", + "(l::'a::field_inverse_zero) / (m * n)", + "((numeral v)::'a::field_inverse_zero) / (numeral w)", + "((numeral v)::'a::field_inverse_zero) / (neg_numeral w)", + "((neg_numeral v)::'a::field_inverse_zero) / (numeral w)", + "((neg_numeral v)::'a::field_inverse_zero) / (neg_numeral w)"], K DivideCancelNumeralFactor.proc)] @@ -678,13 +688,13 @@ val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_right"}, @{thm "divide_Numeral1"}, - @{thm "divide_zero"}, @{thm "divide_Numeral0"}, + @{thm "divide_zero"}, @{thm divide_zero_left}, @{thm "divide_divide_eq_left"}, @{thm "times_divide_eq_left"}, @{thm "times_divide_eq_right"}, @{thm "times_divide_times_eq"}, @{thm "divide_divide_eq_right"}, @{thm "diff_minus"}, @{thm "minus_divide_left"}, - @{thm "Numeral1_eq1_nat"}, @{thm "add_divide_distrib"} RS sym, + @{thm "add_divide_distrib"} RS sym, @{thm field_divide_inverse} RS sym, @{thm inverse_divide}, Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv (Conv.rewr_conv (mk_meta_eq @{thm mult_commute})))) (@{thm field_divide_inverse} RS sym)] @@ -699,8 +709,7 @@ addsimprocs [add_frac_frac_simproc, add_frac_num_simproc, ord_frac_simproc] |> Simplifier.add_cong @{thm "if_weak_cong"}) then_conv - Simplifier.rewrite (HOL_basic_ss addsimps - [@{thm numeral_1_eq_1},@{thm numeral_0_eq_0}] @ @{thms numerals(1-2)}) + Simplifier.rewrite (HOL_basic_ss addsimps [@{thm numeral_1_eq_1}]) end diff -r 29e92b644d6c -r f067afe98049 src/HOL/Tools/numeral_syntax.ML --- a/src/HOL/Tools/numeral_syntax.ML Fri Mar 23 20:32:43 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,93 +0,0 @@ -(* Title: HOL/Tools/numeral_syntax.ML - Authors: Markus Wenzel, TU Muenchen - -Concrete syntax for generic numerals -- preserves leading zeros/ones. -*) - -signature NUMERAL_SYNTAX = -sig - val setup: theory -> theory -end; - -structure Numeral_Syntax: NUMERAL_SYNTAX = -struct - -(* parse translation *) - -local - -fun mk_bin num = - let - fun bit b bs = HOLogic.mk_bit b $ bs; - fun mk 0 = Syntax.const @{const_name Int.Pls} - | mk ~1 = Syntax.const @{const_name Int.Min} - | mk i = let val (q, r) = Integer.div_mod i 2 in bit r (mk q) end; - in mk (#value (Lexicon.read_xnum num)) end; - -in - -fun numeral_tr [(c as Const (@{syntax_const "_constrain"}, _)) $ t $ u] = c $ numeral_tr [t] $ u - | numeral_tr [t as Const (num, _)] = Syntax.const @{const_syntax Int.number_of} $ mk_bin num - | numeral_tr ts = raise TERM ("numeral_tr", ts); - -end; - - -(* print translation *) - -local - -fun dest_bin (Const (@{const_syntax Int.Pls}, _)) = [] - | dest_bin (Const (@{const_syntax Int.Min}, _)) = [~1] - | dest_bin (Const (@{const_syntax Int.Bit0}, _) $ bs) = 0 :: dest_bin bs - | dest_bin (Const (@{const_syntax Int.Bit1}, _) $ bs) = 1 :: dest_bin bs - | dest_bin _ = raise Match; - -fun leading _ [] = 0 - | leading (i: int) (j :: js) = if i = j then 1 + leading i js else 0; - -fun int_of [] = 0 - | int_of (b :: bs) = b + 2 * int_of bs; - -fun dest_bin_str tm = - let - val rev_digs = dest_bin tm; - val (sign, z) = - (case rev rev_digs of - ~1 :: bs => ("-", leading 1 bs) - | bs => ("", leading 0 bs)); - val i = int_of rev_digs; - val num = string_of_int (abs i); - in - if (i = 0 orelse i = 1) andalso z = 0 then raise Match - else sign ^ implode (replicate z "0") ^ num - end; - -fun syntax_numeral t = - Syntax.const @{syntax_const "_Numeral"} $ - (Syntax.const @{syntax_const "_numeral"} $ Syntax.free (dest_bin_str t)); - -in - -fun numeral_tr' ctxt (Type (@{type_name fun}, [_, T])) (t :: ts) = - let val t' = - if not (Config.get ctxt show_types) andalso can Term.dest_Type T then syntax_numeral t - else - Syntax.const @{syntax_const "_constrain"} $ syntax_numeral t $ - Syntax_Phases.term_of_typ ctxt T - in list_comb (t', ts) end - | numeral_tr' _ T (t :: ts) = - if T = dummyT then list_comb (syntax_numeral t, ts) - else raise Match - | numeral_tr' _ _ _ = raise Match; - -end; - - -(* theory setup *) - -val setup = - Sign.add_trfuns ([], [(@{syntax_const "_Numeral"}, numeral_tr)], [], []) #> - Sign.add_advanced_trfunsT [(@{const_syntax Int.number_of}, numeral_tr')]; - -end; diff -r 29e92b644d6c -r f067afe98049 src/HOL/Tools/semiring_normalizer.ML --- a/src/HOL/Tools/semiring_normalizer.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Tools/semiring_normalizer.ML Mon Mar 26 10:56:56 2012 +0200 @@ -179,7 +179,7 @@ (case Rat.quotient_of_rat x of (i, 1) => i | _ => error "int_of_rat: bad int"), conv = fn phi => fn _ => Simplifier.rewrite (HOL_basic_ss addsimps @{thms semiring_norm}) then_conv Simplifier.rewrite (HOL_basic_ss addsimps - (@{thms numeral_1_eq_1} @ @{thms numeral_0_eq_0} @ @{thms numerals(1-2)}))}; + @{thms numeral_1_eq_1})}; fun field_funs key = let @@ -237,13 +237,13 @@ val is_numeral = can dest_numeral; val numeral01_conv = Simplifier.rewrite - (HOL_basic_ss addsimps [@{thm numeral_1_eq_1}, @{thm numeral_0_eq_0}]); + (HOL_basic_ss addsimps [@{thm numeral_1_eq_1}]); val zero1_numeral_conv = - Simplifier.rewrite (HOL_basic_ss addsimps [@{thm numeral_1_eq_1} RS sym, @{thm numeral_0_eq_0} RS sym]); + Simplifier.rewrite (HOL_basic_ss addsimps [@{thm numeral_1_eq_1} RS sym]); fun zerone_conv cv = zero1_numeral_conv then_conv cv then_conv numeral01_conv; -val natarith = [@{thm "add_nat_number_of"}, @{thm "diff_nat_number_of"}, - @{thm "mult_nat_number_of"}, @{thm "eq_nat_number_of"}, - @{thm "less_nat_number_of"}]; +val natarith = [@{thm "numeral_plus_numeral"}, @{thm "diff_nat_numeral"}, + @{thm "numeral_times_numeral"}, @{thm "numeral_eq_iff"}, + @{thm "numeral_less_iff"}]; val nat_add_conv = zerone_conv @@ -251,7 +251,7 @@ (HOL_basic_ss addsimps @{thms arith_simps} @ natarith @ @{thms rel_simps} @ [@{thm if_False}, @{thm if_True}, @{thm Nat.add_0}, @{thm add_Suc}, - @{thm add_number_of_left}, @{thm Suc_eq_plus1}] + @{thm add_numeral_left}, @{thm Suc_eq_plus1}] @ map (fn th => th RS sym) @{thms numerals})); val zeron_tm = @{cterm "0::nat"}; diff -r 29e92b644d6c -r f067afe98049 src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Transcendental.thy Mon Mar 26 10:56:56 2012 +0200 @@ -2044,8 +2044,8 @@ finally show ?thesis by auto qed -lemma tan_periodic_n[simp]: "tan (x + number_of n * pi) = tan x" - using tan_periodic_int[of _ "number_of n" ] unfolding real_number_of . +lemma tan_periodic_n[simp]: "tan (x + numeral n * pi) = tan x" + using tan_periodic_int[of _ "numeral n" ] unfolding real_numeral . subsection {* Inverse Trigonometric Functions *} diff -r 29e92b644d6c -r f067afe98049 src/HOL/Unix/Unix.thy --- a/src/HOL/Unix/Unix.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Unix/Unix.thy Mon Mar 26 10:56:56 2012 +0200 @@ -843,7 +843,9 @@ neither owned nor writable by @{term user\<^isub>1}. *} -definition + + +definition invariant where "invariant root path = (\att dir. access root path user\<^isub>1 {} = Some (Env att dir) \ dir \ empty \ diff -r 29e92b644d6c -r f067afe98049 src/HOL/Word/Bit_Int.thy --- a/src/HOL/Word/Bit_Int.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Word/Bit_Int.thy Mon Mar 26 10:56:56 2012 +0200 @@ -50,11 +50,13 @@ unfolding int_not_def Bit_def by (cases b, simp_all) lemma int_not_simps [simp]: - "NOT Int.Pls = Int.Min" - "NOT Int.Min = Int.Pls" - "NOT (Int.Bit0 w) = Int.Bit1 (NOT w)" - "NOT (Int.Bit1 w) = Int.Bit0 (NOT w)" - unfolding int_not_def Pls_def Min_def Bit0_def Bit1_def by simp_all + "NOT (0::int) = -1" + "NOT (1::int) = -2" + "NOT (-1::int) = 0" + "NOT (numeral w::int) = neg_numeral (w + Num.One)" + "NOT (neg_numeral (Num.Bit0 w)::int) = numeral (Num.BitM w)" + "NOT (neg_numeral (Num.Bit1 w)::int) = numeral (Num.Bit0 w)" + unfolding int_not_def by simp_all lemma int_not_not [simp]: "NOT (NOT (x::int)) = x" unfolding int_not_def by simp @@ -65,12 +67,6 @@ lemma int_and_m1 [simp]: "(-1::int) AND x = x" by (simp add: bitAND_int.simps) -lemma int_and_Pls [simp]: "Int.Pls AND x = Int.Pls" - unfolding Pls_def by simp - -lemma int_and_Min [simp]: "Int.Min AND x = x" - unfolding Min_def by simp - lemma Bit_eq_0_iff: "w BIT b = 0 \ w = 0 \ b = 0" by (subst BIT_eq_iff [symmetric], simp) @@ -81,17 +77,10 @@ "(x BIT b) AND (y BIT c) = (x AND y) BIT (b AND c)" by (subst bitAND_int.simps, simp add: Bit_eq_0_iff Bit_eq_m1_iff) -lemma int_and_Bits2 [simp]: - "(Int.Bit0 x) AND (Int.Bit0 y) = Int.Bit0 (x AND y)" - "(Int.Bit0 x) AND (Int.Bit1 y) = Int.Bit0 (x AND y)" - "(Int.Bit1 x) AND (Int.Bit0 y) = Int.Bit0 (x AND y)" - "(Int.Bit1 x) AND (Int.Bit1 y) = Int.Bit1 (x AND y)" - unfolding BIT_simps [symmetric] int_and_Bits by simp_all - -lemma int_or_Pls [simp]: "Int.Pls OR x = x" +lemma int_or_zero [simp]: "(0::int) OR x = x" unfolding int_or_def by simp -lemma int_or_Min [simp]: "Int.Min OR x = Int.Min" +lemma int_or_minus1 [simp]: "(-1::int) OR x = -1" unfolding int_or_def by simp lemma bit_or_def: "(b::bit) OR c = NOT (NOT b AND NOT c)" @@ -101,14 +90,7 @@ "(x BIT b) OR (y BIT c) = (x OR y) BIT (b OR c)" unfolding int_or_def bit_or_def by simp -lemma int_or_Bits2 [simp]: - "(Int.Bit0 x) OR (Int.Bit0 y) = Int.Bit0 (x OR y)" - "(Int.Bit0 x) OR (Int.Bit1 y) = Int.Bit1 (x OR y)" - "(Int.Bit1 x) OR (Int.Bit0 y) = Int.Bit1 (x OR y)" - "(Int.Bit1 x) OR (Int.Bit1 y) = Int.Bit1 (x OR y)" - unfolding int_or_def by simp_all - -lemma int_xor_Pls [simp]: "Int.Pls XOR x = x" +lemma int_xor_zero [simp]: "(0::int) XOR x = x" unfolding int_xor_def by simp lemma bit_xor_def: "(b::bit) XOR c = (b AND NOT c) OR (NOT b AND c)" @@ -118,13 +100,6 @@ "(x BIT b) XOR (y BIT c) = (x XOR y) BIT (b XOR c)" unfolding int_xor_def bit_xor_def by simp -lemma int_xor_Bits2 [simp]: - "(Int.Bit0 x) XOR (Int.Bit0 y) = Int.Bit0 (x XOR y)" - "(Int.Bit0 x) XOR (Int.Bit1 y) = Int.Bit1 (x XOR y)" - "(Int.Bit1 x) XOR (Int.Bit0 y) = Int.Bit1 (x XOR y)" - "(Int.Bit1 x) XOR (Int.Bit1 y) = Int.Bit0 (x XOR y)" - unfolding BIT_simps [symmetric] int_xor_Bits by simp_all - subsubsection {* Binary destructors *} lemma bin_rest_NOT [simp]: "bin_rest (NOT x) = NOT (bin_rest x)" @@ -166,22 +141,22 @@ subsubsection {* Derived properties *} -lemma int_xor_Min [simp]: "Int.Min XOR x = NOT x" +lemma int_xor_minus1 [simp]: "(-1::int) XOR x = NOT x" by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_xor_extra_simps [simp]: - "w XOR Int.Pls = w" - "w XOR Int.Min = NOT w" + "w XOR (0::int) = w" + "w XOR (-1::int) = NOT w" by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_or_extra_simps [simp]: - "w OR Int.Pls = w" - "w OR Int.Min = Int.Min" + "w OR (0::int) = w" + "w OR (-1::int) = -1" by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_and_extra_simps [simp]: - "w AND Int.Pls = Int.Pls" - "w AND Int.Min = w" + "w AND (0::int) = 0" + "w AND (-1::int) = w" by (auto simp add: bin_eq_iff bin_nth_ops) (* commutativity of the above *) @@ -195,12 +170,12 @@ lemma bin_ops_same [simp]: "(x::int) AND x = x" "(x::int) OR x = x" - "(x::int) XOR x = Int.Pls" + "(x::int) XOR x = 0" by (auto simp add: bin_eq_iff bin_nth_ops) lemmas bin_log_esimps = int_and_extra_simps int_or_extra_simps int_xor_extra_simps - int_and_Pls int_and_Min int_or_Pls int_or_Min int_xor_Pls int_xor_Min + int_and_0 int_and_m1 int_or_zero int_or_minus1 int_xor_zero int_xor_minus1 (* basic properties of logical (bit-wise) operations *) @@ -262,6 +237,106 @@ declare bin_ops_comm [simp] bbw_assocs [simp] *) +subsubsection {* Simplification with numerals *} + +text {* Cases for @{text "0"} and @{text "-1"} are already covered by + other simp rules. *} + +lemma bin_rl_eqI: "\bin_rest x = bin_rest y; bin_last x = bin_last y\ \ x = y" + by (metis bin_rl_simp) + +lemma bin_rest_neg_numeral_BitM [simp]: + "bin_rest (neg_numeral (Num.BitM w)) = neg_numeral w" + by (simp only: BIT_bin_simps [symmetric] bin_rest_BIT) + +lemma bin_last_neg_numeral_BitM [simp]: + "bin_last (neg_numeral (Num.BitM w)) = 1" + by (simp only: BIT_bin_simps [symmetric] bin_last_BIT) + +text {* FIXME: The rule sets below are very large (24 rules for each + operator). Is there a simpler way to do this? *} + +lemma int_and_numerals [simp]: + "numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (numeral x AND numeral y) BIT 0" + "numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (numeral x AND numeral y) BIT 0" + "numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (numeral x AND numeral y) BIT 0" + "numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = (numeral x AND numeral y) BIT 1" + "numeral (Num.Bit0 x) AND neg_numeral (Num.Bit0 y) = (numeral x AND neg_numeral y) BIT 0" + "numeral (Num.Bit0 x) AND neg_numeral (Num.Bit1 y) = (numeral x AND neg_numeral (y + Num.One)) BIT 0" + "numeral (Num.Bit1 x) AND neg_numeral (Num.Bit0 y) = (numeral x AND neg_numeral y) BIT 0" + "numeral (Num.Bit1 x) AND neg_numeral (Num.Bit1 y) = (numeral x AND neg_numeral (y + Num.One)) BIT 1" + "neg_numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (neg_numeral x AND numeral y) BIT 0" + "neg_numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (neg_numeral x AND numeral y) BIT 0" + "neg_numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (neg_numeral (x + Num.One) AND numeral y) BIT 0" + "neg_numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = (neg_numeral (x + Num.One) AND numeral y) BIT 1" + "neg_numeral (Num.Bit0 x) AND neg_numeral (Num.Bit0 y) = (neg_numeral x AND neg_numeral y) BIT 0" + "neg_numeral (Num.Bit0 x) AND neg_numeral (Num.Bit1 y) = (neg_numeral x AND neg_numeral (y + Num.One)) BIT 0" + "neg_numeral (Num.Bit1 x) AND neg_numeral (Num.Bit0 y) = (neg_numeral (x + Num.One) AND neg_numeral y) BIT 0" + "neg_numeral (Num.Bit1 x) AND neg_numeral (Num.Bit1 y) = (neg_numeral (x + Num.One) AND neg_numeral (y + Num.One)) BIT 1" + "(1::int) AND numeral (Num.Bit0 y) = 0" + "(1::int) AND numeral (Num.Bit1 y) = 1" + "(1::int) AND neg_numeral (Num.Bit0 y) = 0" + "(1::int) AND neg_numeral (Num.Bit1 y) = 1" + "numeral (Num.Bit0 x) AND (1::int) = 0" + "numeral (Num.Bit1 x) AND (1::int) = 1" + "neg_numeral (Num.Bit0 x) AND (1::int) = 0" + "neg_numeral (Num.Bit1 x) AND (1::int) = 1" + by (rule bin_rl_eqI, simp, simp)+ + +lemma int_or_numerals [simp]: + "numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (numeral x OR numeral y) BIT 0" + "numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = (numeral x OR numeral y) BIT 1" + "numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = (numeral x OR numeral y) BIT 1" + "numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = (numeral x OR numeral y) BIT 1" + "numeral (Num.Bit0 x) OR neg_numeral (Num.Bit0 y) = (numeral x OR neg_numeral y) BIT 0" + "numeral (Num.Bit0 x) OR neg_numeral (Num.Bit1 y) = (numeral x OR neg_numeral (y + Num.One)) BIT 1" + "numeral (Num.Bit1 x) OR neg_numeral (Num.Bit0 y) = (numeral x OR neg_numeral y) BIT 1" + "numeral (Num.Bit1 x) OR neg_numeral (Num.Bit1 y) = (numeral x OR neg_numeral (y + Num.One)) BIT 1" + "neg_numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (neg_numeral x OR numeral y) BIT 0" + "neg_numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = (neg_numeral x OR numeral y) BIT 1" + "neg_numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = (neg_numeral (x + Num.One) OR numeral y) BIT 1" + "neg_numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = (neg_numeral (x + Num.One) OR numeral y) BIT 1" + "neg_numeral (Num.Bit0 x) OR neg_numeral (Num.Bit0 y) = (neg_numeral x OR neg_numeral y) BIT 0" + "neg_numeral (Num.Bit0 x) OR neg_numeral (Num.Bit1 y) = (neg_numeral x OR neg_numeral (y + Num.One)) BIT 1" + "neg_numeral (Num.Bit1 x) OR neg_numeral (Num.Bit0 y) = (neg_numeral (x + Num.One) OR neg_numeral y) BIT 1" + "neg_numeral (Num.Bit1 x) OR neg_numeral (Num.Bit1 y) = (neg_numeral (x + Num.One) OR neg_numeral (y + Num.One)) BIT 1" + "(1::int) OR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)" + "(1::int) OR numeral (Num.Bit1 y) = numeral (Num.Bit1 y)" + "(1::int) OR neg_numeral (Num.Bit0 y) = neg_numeral (Num.BitM y)" + "(1::int) OR neg_numeral (Num.Bit1 y) = neg_numeral (Num.Bit1 y)" + "numeral (Num.Bit0 x) OR (1::int) = numeral (Num.Bit1 x)" + "numeral (Num.Bit1 x) OR (1::int) = numeral (Num.Bit1 x)" + "neg_numeral (Num.Bit0 x) OR (1::int) = neg_numeral (Num.BitM x)" + "neg_numeral (Num.Bit1 x) OR (1::int) = neg_numeral (Num.Bit1 x)" + by (rule bin_rl_eqI, simp, simp)+ + +lemma int_xor_numerals [simp]: + "numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (numeral x XOR numeral y) BIT 0" + "numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = (numeral x XOR numeral y) BIT 1" + "numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = (numeral x XOR numeral y) BIT 1" + "numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (numeral x XOR numeral y) BIT 0" + "numeral (Num.Bit0 x) XOR neg_numeral (Num.Bit0 y) = (numeral x XOR neg_numeral y) BIT 0" + "numeral (Num.Bit0 x) XOR neg_numeral (Num.Bit1 y) = (numeral x XOR neg_numeral (y + Num.One)) BIT 1" + "numeral (Num.Bit1 x) XOR neg_numeral (Num.Bit0 y) = (numeral x XOR neg_numeral y) BIT 1" + "numeral (Num.Bit1 x) XOR neg_numeral (Num.Bit1 y) = (numeral x XOR neg_numeral (y + Num.One)) BIT 0" + "neg_numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (neg_numeral x XOR numeral y) BIT 0" + "neg_numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = (neg_numeral x XOR numeral y) BIT 1" + "neg_numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = (neg_numeral (x + Num.One) XOR numeral y) BIT 1" + "neg_numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (neg_numeral (x + Num.One) XOR numeral y) BIT 0" + "neg_numeral (Num.Bit0 x) XOR neg_numeral (Num.Bit0 y) = (neg_numeral x XOR neg_numeral y) BIT 0" + "neg_numeral (Num.Bit0 x) XOR neg_numeral (Num.Bit1 y) = (neg_numeral x XOR neg_numeral (y + Num.One)) BIT 1" + "neg_numeral (Num.Bit1 x) XOR neg_numeral (Num.Bit0 y) = (neg_numeral (x + Num.One) XOR neg_numeral y) BIT 1" + "neg_numeral (Num.Bit1 x) XOR neg_numeral (Num.Bit1 y) = (neg_numeral (x + Num.One) XOR neg_numeral (y + Num.One)) BIT 0" + "(1::int) XOR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)" + "(1::int) XOR numeral (Num.Bit1 y) = numeral (Num.Bit0 y)" + "(1::int) XOR neg_numeral (Num.Bit0 y) = neg_numeral (Num.BitM y)" + "(1::int) XOR neg_numeral (Num.Bit1 y) = neg_numeral (Num.Bit0 (y + Num.One))" + "numeral (Num.Bit0 x) XOR (1::int) = numeral (Num.Bit1 x)" + "numeral (Num.Bit1 x) XOR (1::int) = numeral (Num.Bit0 x)" + "neg_numeral (Num.Bit0 x) XOR (1::int) = neg_numeral (Num.BitM x)" + "neg_numeral (Num.Bit1 x) XOR (1::int) = neg_numeral (Num.Bit0 (x + Num.One))" + by (rule bin_rl_eqI, simp, simp)+ + subsubsection {* Interactions with arithmetic *} lemma plus_and_or [rule_format]: @@ -282,7 +357,6 @@ "bin_sign (y::int) = 0 ==> x <= x OR y" apply (induct y arbitrary: x rule: bin_induct) apply clarsimp - apply (simp only: Min_def) apply clarsimp apply (case_tac x rule: bin_exhaust) apply (case_tac b) @@ -293,13 +367,20 @@ lemmas int_and_le = xtr3 [OF bbw_ao_absorbs (2) [THEN conjunct2, symmetric] le_int_or] +lemma add_BIT_simps [simp]: (* FIXME: move *) + "x BIT 0 + y BIT 0 = (x + y) BIT 0" + "x BIT 0 + y BIT 1 = (x + y) BIT 1" + "x BIT 1 + y BIT 0 = (x + y) BIT 1" + "x BIT 1 + y BIT 1 = (x + y + 1) BIT 0" + by (simp_all add: Bit_B0_2t Bit_B1_2t) + (* interaction between bit-wise and arithmetic *) (* good example of bin_induction *) -lemma bin_add_not: "x + NOT x = Int.Min" +lemma bin_add_not: "x + NOT x = (-1::int)" apply (induct x rule: bin_induct) apply clarsimp apply clarsimp - apply (case_tac bit, auto simp: BIT_simps) + apply (case_tac bit, auto) done subsubsection {* Truncating results of bit-wise operations *} @@ -418,8 +499,10 @@ lemmas bin_sc_Suc_minus = trans [OF bin_sc_minus [symmetric] bin_sc.Suc] -lemmas bin_sc_Suc_pred [simp] = - bin_sc_Suc_minus [of "number_of bin", simplified nobm1] for bin +lemma bin_sc_numeral [simp]: + "bin_sc (numeral k) b w = + bin_sc (numeral k - 1) b (bin_rest w) BIT bin_last w" + by (subst expand_Suc, rule bin_sc.Suc) subsection {* Splitting and concatenation *} diff -r 29e92b644d6c -r f067afe98049 src/HOL/Word/Bit_Representation.thy --- a/src/HOL/Word/Bit_Representation.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Word/Bit_Representation.thy Mon Mar 26 10:56:56 2012 +0200 @@ -47,41 +47,49 @@ by (metis bin_rest_BIT bin_last_BIT) lemma BIT_bin_simps [simp]: - "number_of w BIT 0 = number_of (Int.Bit0 w)" - "number_of w BIT 1 = number_of (Int.Bit1 w)" - unfolding Bit_def number_of_is_id numeral_simps by simp_all + "numeral k BIT 0 = numeral (Num.Bit0 k)" + "numeral k BIT 1 = numeral (Num.Bit1 k)" + "neg_numeral k BIT 0 = neg_numeral (Num.Bit0 k)" + "neg_numeral k BIT 1 = neg_numeral (Num.BitM k)" + unfolding neg_numeral_def numeral.simps numeral_BitM + unfolding Bit_def bitval_simps + by (simp_all del: arith_simps add_numeral_special diff_numeral_special) lemma BIT_special_simps [simp]: shows "0 BIT 0 = 0" and "0 BIT 1 = 1" and "1 BIT 0 = 2" and "1 BIT 1 = 3" unfolding Bit_def by simp_all +lemma BitM_inc: "Num.BitM (Num.inc w) = Num.Bit1 w" + by (induct w, simp_all) + +lemma expand_BIT: + "numeral (Num.Bit0 w) = numeral w BIT 0" + "numeral (Num.Bit1 w) = numeral w BIT 1" + "neg_numeral (Num.Bit0 w) = neg_numeral w BIT 0" + "neg_numeral (Num.Bit1 w) = neg_numeral (w + Num.One) BIT 1" + unfolding add_One by (simp_all add: BitM_inc) + lemma bin_last_numeral_simps [simp]: "bin_last 0 = 0" "bin_last 1 = 1" "bin_last -1 = 1" - "bin_last (number_of (Int.Bit0 w)) = 0" - "bin_last (number_of (Int.Bit1 w)) = 1" - unfolding bin_last_def by simp_all + "bin_last Numeral1 = 1" + "bin_last (numeral (Num.Bit0 w)) = 0" + "bin_last (numeral (Num.Bit1 w)) = 1" + "bin_last (neg_numeral (Num.Bit0 w)) = 0" + "bin_last (neg_numeral (Num.Bit1 w)) = 1" + unfolding expand_BIT bin_last_BIT by (simp_all add: bin_last_def) lemma bin_rest_numeral_simps [simp]: "bin_rest 0 = 0" "bin_rest 1 = 0" "bin_rest -1 = -1" - "bin_rest (number_of (Int.Bit0 w)) = number_of w" - "bin_rest (number_of (Int.Bit1 w)) = number_of w" - unfolding bin_rest_def by simp_all - -lemma BIT_B0_eq_Bit0: "w BIT 0 = Int.Bit0 w" - unfolding Bit_def Bit0_def by simp - -lemma BIT_B1_eq_Bit1: "w BIT 1 = Int.Bit1 w" - unfolding Bit_def Bit1_def by simp - -lemmas BIT_simps = BIT_B0_eq_Bit0 BIT_B1_eq_Bit1 - -lemma number_of_False_cong: - "False \ number_of x = number_of y" - by (rule FalseE) + "bin_rest Numeral1 = 0" + "bin_rest (numeral (Num.Bit0 w)) = numeral w" + "bin_rest (numeral (Num.Bit1 w)) = numeral w" + "bin_rest (neg_numeral (Num.Bit0 w)) = neg_numeral w" + "bin_rest (neg_numeral (Num.Bit1 w)) = neg_numeral (w + Num.One)" + unfolding expand_BIT bin_rest_BIT by (simp_all add: bin_rest_def) lemma less_Bits: "(v BIT b < w BIT c) = (v < w | v <= w & b = (0::bit) & c = (1::bit))" @@ -121,11 +129,7 @@ done lemma bin_ex_rl: "EX w b. w BIT b = bin" - apply (unfold Bit_def) - apply (cases "even bin") - apply (clarsimp simp: even_equiv_def) - apply (auto simp: odd_equiv_def bitval_def split: bit.split) - done + by (metis bin_rl_simp) lemma bin_exhaust: assumes Q: "\x b. bin = x BIT b \ Q" @@ -144,18 +148,18 @@ | Suc: "bin_nth w (Suc n) = bin_nth (bin_rest w) n" lemma bin_abs_lem: - "bin = (w BIT b) ==> ~ bin = Int.Min --> ~ bin = Int.Pls --> + "bin = (w BIT b) ==> bin ~= -1 --> bin ~= 0 --> nat (abs w) < nat (abs bin)" apply clarsimp - apply (unfold Pls_def Min_def Bit_def) + apply (unfold Bit_def) apply (cases b) apply (clarsimp, arith) apply (clarsimp, arith) done lemma bin_induct: - assumes PPls: "P Int.Pls" - and PMin: "P Int.Min" + assumes PPls: "P 0" + and PMin: "P -1" and PBit: "!!bin bit. P bin ==> P (bin BIT bit)" shows "P bin" apply (rule_tac P=P and a=bin and f1="nat o abs" @@ -166,54 +170,22 @@ apply (auto simp add : PPls PMin PBit) done -lemma numeral_induct: - assumes Pls: "P Int.Pls" - assumes Min: "P Int.Min" - assumes Bit0: "\w. \P w; w \ Int.Pls\ \ P (Int.Bit0 w)" - assumes Bit1: "\w. \P w; w \ Int.Min\ \ P (Int.Bit1 w)" - shows "P x" - apply (induct x rule: bin_induct) - apply (rule Pls) - apply (rule Min) - apply (case_tac bit) - apply (case_tac "bin = Int.Pls") - apply (simp add: BIT_simps) - apply (simp add: Bit0 BIT_simps) - apply (case_tac "bin = Int.Min") - apply (simp add: BIT_simps) - apply (simp add: Bit1 BIT_simps) - done - -lemma bin_rest_simps [simp]: - "bin_rest Int.Pls = Int.Pls" - "bin_rest Int.Min = Int.Min" - "bin_rest (Int.Bit0 w) = w" - "bin_rest (Int.Bit1 w) = w" - unfolding numeral_simps by (auto simp: bin_rest_def) - -lemma bin_last_simps [simp]: - "bin_last Int.Pls = (0::bit)" - "bin_last Int.Min = (1::bit)" - "bin_last (Int.Bit0 w) = (0::bit)" - "bin_last (Int.Bit1 w) = (1::bit)" - unfolding numeral_simps by (auto simp: bin_last_def z1pmod2) - lemma Bit_div2 [simp]: "(w BIT b) div 2 = w" unfolding bin_rest_def [symmetric] by (rule bin_rest_BIT) lemma bin_nth_lem [rule_format]: "ALL y. bin_nth x = bin_nth y --> x = y" - apply (induct x rule: bin_induct [unfolded Pls_def Min_def]) + apply (induct x rule: bin_induct) apply safe apply (erule rev_mp) - apply (induct_tac y rule: bin_induct [unfolded Pls_def Min_def]) + apply (induct_tac y rule: bin_induct) apply safe apply (drule_tac x=0 in fun_cong, force) apply (erule notE, rule ext, drule_tac x="Suc x" in fun_cong, force) apply (drule_tac x=0 in fun_cong, force) apply (erule rev_mp) - apply (induct_tac y rule: bin_induct [unfolded Pls_def Min_def]) + apply (induct_tac y rule: bin_induct) apply safe apply (drule_tac x=0 in fun_cong, force) apply (erule notE, rule ext, @@ -244,15 +216,9 @@ lemma bin_nth_1 [simp]: "bin_nth 1 n \ n = 0" by (cases n) simp_all -lemma bin_nth_Pls [simp]: "~ bin_nth Int.Pls n" - by (induct n) auto (* FIXME: delete *) - lemma bin_nth_minus1 [simp]: "bin_nth -1 n" by (induct n) auto -lemma bin_nth_Min [simp]: "bin_nth Int.Min n" - by (induct n) auto (* FIXME: delete *) - lemma bin_nth_0_BIT: "bin_nth (w BIT b) 0 = (b = (1::bit))" by auto @@ -262,20 +228,20 @@ lemma bin_nth_minus [simp]: "0 < n ==> bin_nth (w BIT b) n = bin_nth w (n - 1)" by (cases n) auto -lemma bin_nth_minus_Bit0 [simp]: - "0 < n ==> bin_nth (number_of (Int.Bit0 w)) n = bin_nth (number_of w) (n - 1)" - using bin_nth_minus [where w="number_of w" and b="(0::bit)"] by simp +lemma bin_nth_numeral: + "bin_rest x = y \ bin_nth x (numeral n) = bin_nth y (numeral n - 1)" + by (subst expand_Suc, simp only: bin_nth.simps) -lemma bin_nth_minus_Bit1 [simp]: - "0 < n ==> bin_nth (number_of (Int.Bit1 w)) n = bin_nth (number_of w) (n - 1)" - using bin_nth_minus [where w="number_of w" and b="(1::bit)"] by simp - -lemmas bin_nth_0 = bin_nth.simps(1) -lemmas bin_nth_Suc = bin_nth.simps(2) +lemmas bin_nth_numeral_simps [simp] = + bin_nth_numeral [OF bin_rest_numeral_simps(2)] + bin_nth_numeral [OF bin_rest_numeral_simps(5)] + bin_nth_numeral [OF bin_rest_numeral_simps(6)] + bin_nth_numeral [OF bin_rest_numeral_simps(7)] + bin_nth_numeral [OF bin_rest_numeral_simps(8)] lemmas bin_nth_simps = - bin_nth_0 bin_nth_Suc bin_nth_zero bin_nth_minus1 bin_nth_minus - bin_nth_minus_Bit0 bin_nth_minus_Bit1 + bin_nth.Z bin_nth.Suc bin_nth_zero bin_nth_minus1 + bin_nth_numeral_simps subsection {* Truncating binary integers *} @@ -286,9 +252,8 @@ lemma bin_sign_simps [simp]: "bin_sign 0 = 0" "bin_sign 1 = 0" - "bin_sign -1 = -1" - "bin_sign (number_of (Int.Bit0 w)) = bin_sign (number_of w)" - "bin_sign (number_of (Int.Bit1 w)) = bin_sign (number_of w)" + "bin_sign (numeral k) = 0" + "bin_sign (neg_numeral k) = -1" "bin_sign (w BIT b) = bin_sign w" unfolding bin_sign_def Bit_def bitval_def by (simp_all split: bit.split) @@ -309,17 +274,15 @@ by (induct n arbitrary: w) auto lemma bintrunc_mod2p: "bintrunc n w = (w mod 2 ^ n)" - apply (induct n arbitrary: w) - apply simp + apply (induct n arbitrary: w, clarsimp) apply (simp add: bin_last_def bin_rest_def Bit_def zmod_zmult2_eq) done lemma sbintrunc_mod2p: "sbintrunc n w = (w + 2 ^ n) mod 2 ^ (Suc n) - 2 ^ n" apply (induct n arbitrary: w) - apply clarsimp + apply simp apply (subst mod_add_left_eq) apply (simp add: bin_last_def) - apply simp apply (simp add: bin_last_def bin_rest_def Bit_def) apply (clarsimp simp: mod_mult_mult1 [symmetric] zmod_zdiv_equality [THEN diff_eq_eq [THEN iffD2 [THEN sym]]]) @@ -342,20 +305,32 @@ lemma bintrunc_Suc_numeral: "bintrunc (Suc n) 1 = 1" "bintrunc (Suc n) -1 = bintrunc n -1 BIT 1" - "bintrunc (Suc n) (number_of (Int.Bit0 w)) = bintrunc n (number_of w) BIT 0" - "bintrunc (Suc n) (number_of (Int.Bit1 w)) = bintrunc n (number_of w) BIT 1" + "bintrunc (Suc n) (numeral (Num.Bit0 w)) = bintrunc n (numeral w) BIT 0" + "bintrunc (Suc n) (numeral (Num.Bit1 w)) = bintrunc n (numeral w) BIT 1" + "bintrunc (Suc n) (neg_numeral (Num.Bit0 w)) = + bintrunc n (neg_numeral w) BIT 0" + "bintrunc (Suc n) (neg_numeral (Num.Bit1 w)) = + bintrunc n (neg_numeral (w + Num.One)) BIT 1" by simp_all lemma sbintrunc_0_numeral [simp]: "sbintrunc 0 1 = -1" - "sbintrunc 0 (number_of (Int.Bit0 w)) = 0" - "sbintrunc 0 (number_of (Int.Bit1 w)) = -1" + "sbintrunc 0 (numeral (Num.Bit0 w)) = 0" + "sbintrunc 0 (numeral (Num.Bit1 w)) = -1" + "sbintrunc 0 (neg_numeral (Num.Bit0 w)) = 0" + "sbintrunc 0 (neg_numeral (Num.Bit1 w)) = -1" by simp_all lemma sbintrunc_Suc_numeral: "sbintrunc (Suc n) 1 = 1" - "sbintrunc (Suc n) (number_of (Int.Bit0 w)) = sbintrunc n (number_of w) BIT 0" - "sbintrunc (Suc n) (number_of (Int.Bit1 w)) = sbintrunc n (number_of w) BIT 1" + "sbintrunc (Suc n) (numeral (Num.Bit0 w)) = + sbintrunc n (numeral w) BIT 0" + "sbintrunc (Suc n) (numeral (Num.Bit1 w)) = + sbintrunc n (numeral w) BIT 1" + "sbintrunc (Suc n) (neg_numeral (Num.Bit0 w)) = + sbintrunc n (neg_numeral w) BIT 0" + "sbintrunc (Suc n) (neg_numeral (Num.Bit1 w)) = + sbintrunc n (neg_numeral (w + Num.One)) BIT 1" by simp_all lemma bit_bool: @@ -366,7 +341,7 @@ lemma bin_sign_lem: "(bin_sign (sbintrunc n bin) = -1) = bin_nth bin n" apply (induct n arbitrary: bin) - apply (case_tac bin rule: bin_exhaust, case_tac b, auto)+ + apply (case_tac bin rule: bin_exhaust, case_tac b, auto) done lemma nth_bintr: "bin_nth (bintrunc m w) n = (n < m & bin_nth w n)" @@ -388,14 +363,14 @@ by (cases n) auto lemma bin_nth_Bit0: - "bin_nth (number_of (Int.Bit0 w)) n \ - (\m. n = Suc m \ bin_nth (number_of w) m)" - using bin_nth_Bit [where w="number_of w" and b="(0::bit)"] by simp + "bin_nth (numeral (Num.Bit0 w)) n \ + (\m. n = Suc m \ bin_nth (numeral w) m)" + using bin_nth_Bit [where w="numeral w" and b="(0::bit)"] by simp lemma bin_nth_Bit1: - "bin_nth (number_of (Int.Bit1 w)) n \ - n = 0 \ (\m. n = Suc m \ bin_nth (number_of w) m)" - using bin_nth_Bit [where w="number_of w" and b="(1::bit)"] by simp + "bin_nth (numeral (Num.Bit1 w)) n \ + n = 0 \ (\m. n = Suc m \ bin_nth (numeral w) m)" + using bin_nth_Bit [where w="numeral w" and b="(1::bit)"] by simp lemma bintrunc_bintrunc_l: "n <= m ==> (bintrunc m (bintrunc n w) = bintrunc n w)" @@ -422,72 +397,47 @@ done lemmas bintrunc_Pls = - bintrunc.Suc [where bin="Int.Pls", simplified bin_last_simps bin_rest_simps] + bintrunc.Suc [where bin="0", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas bintrunc_Min [simp] = - bintrunc.Suc [where bin="Int.Min", simplified bin_last_simps bin_rest_simps] + bintrunc.Suc [where bin="-1", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas bintrunc_BIT [simp] = bintrunc.Suc [where bin="w BIT b", simplified bin_last_BIT bin_rest_BIT] for w b -lemma bintrunc_Bit0 [simp]: - "bintrunc (Suc n) (Int.Bit0 w) = Int.Bit0 (bintrunc n w)" - using bintrunc_BIT [where b="(0::bit)"] by (simp add: BIT_simps) - -lemma bintrunc_Bit1 [simp]: - "bintrunc (Suc n) (Int.Bit1 w) = Int.Bit1 (bintrunc n w)" - using bintrunc_BIT [where b="(1::bit)"] by (simp add: BIT_simps) - lemmas bintrunc_Sucs = bintrunc_Pls bintrunc_Min bintrunc_BIT - bintrunc_Bit0 bintrunc_Bit1 bintrunc_Suc_numeral lemmas sbintrunc_Suc_Pls = - sbintrunc.Suc [where bin="Int.Pls", simplified bin_last_simps bin_rest_simps] + sbintrunc.Suc [where bin="0", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Suc_Min = - sbintrunc.Suc [where bin="Int.Min", simplified bin_last_simps bin_rest_simps] + sbintrunc.Suc [where bin="-1", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Suc_BIT [simp] = sbintrunc.Suc [where bin="w BIT b", simplified bin_last_BIT bin_rest_BIT] for w b -lemma sbintrunc_Suc_Bit0 [simp]: - "sbintrunc (Suc n) (Int.Bit0 w) = Int.Bit0 (sbintrunc n w)" - using sbintrunc_Suc_BIT [where b="(0::bit)"] by (simp add: BIT_simps) - -lemma sbintrunc_Suc_Bit1 [simp]: - "sbintrunc (Suc n) (Int.Bit1 w) = Int.Bit1 (sbintrunc n w)" - using sbintrunc_Suc_BIT [where b="(1::bit)"] by (simp add: BIT_simps) - lemmas sbintrunc_Sucs = sbintrunc_Suc_Pls sbintrunc_Suc_Min sbintrunc_Suc_BIT - sbintrunc_Suc_Bit0 sbintrunc_Suc_Bit1 sbintrunc_Suc_numeral lemmas sbintrunc_Pls = - sbintrunc.Z [where bin="Int.Pls", - simplified bin_last_simps bin_rest_simps bit.simps] + sbintrunc.Z [where bin="0", + simplified bin_last_numeral_simps bin_rest_numeral_simps bit.simps] lemmas sbintrunc_Min = - sbintrunc.Z [where bin="Int.Min", - simplified bin_last_simps bin_rest_simps bit.simps] + sbintrunc.Z [where bin="-1", + simplified bin_last_numeral_simps bin_rest_numeral_simps bit.simps] lemmas sbintrunc_0_BIT_B0 [simp] = sbintrunc.Z [where bin="w BIT (0::bit)", - simplified bin_last_simps bin_rest_simps bit.simps] for w + simplified bin_last_numeral_simps bin_rest_numeral_simps bit.simps] for w lemmas sbintrunc_0_BIT_B1 [simp] = sbintrunc.Z [where bin="w BIT (1::bit)", - simplified bin_last_simps bin_rest_simps bit.simps] for w - -lemma sbintrunc_0_Bit0 [simp]: "sbintrunc 0 (Int.Bit0 w) = 0" - using sbintrunc_0_BIT_B0 by simp - -lemma sbintrunc_0_Bit1 [simp]: "sbintrunc 0 (Int.Bit1 w) = -1" - using sbintrunc_0_BIT_B1 by simp + simplified bin_last_BIT bin_rest_numeral_simps bit.simps] for w lemmas sbintrunc_0_simps = sbintrunc_Pls sbintrunc_Min sbintrunc_0_BIT_B0 sbintrunc_0_BIT_B1 - sbintrunc_0_Bit0 sbintrunc_0_Bit1 lemmas bintrunc_simps = bintrunc.Z bintrunc_Sucs lemmas sbintrunc_simps = sbintrunc_0_simps sbintrunc_Sucs @@ -505,15 +455,6 @@ lemmas sbintrunc_minus_simps = sbintrunc_Sucs [THEN [2] sbintrunc_minus [symmetric, THEN trans]] -lemma bintrunc_n_Pls [simp]: - "bintrunc n Int.Pls = Int.Pls" - unfolding Pls_def by simp - -lemma sbintrunc_n_PM [simp]: - "sbintrunc n Int.Pls = Int.Pls" - "sbintrunc n Int.Min = Int.Min" - unfolding Pls_def Min_def by simp_all - lemmas thobini1 = arg_cong [where f = "%w. w BIT b"] for b lemmas bintrunc_BIT_I = trans [OF bintrunc_BIT thobini1] @@ -600,15 +541,39 @@ lemmas nat_non0_gr = trans [OF iszero_def [THEN Not_eq_iff [THEN iffD2]] refl] -lemmas bintrunc_pred_simps [simp] = - bintrunc_minus_simps [of "number_of bin", simplified nobm1] for bin +lemma bintrunc_numeral: + "bintrunc (numeral k) x = + bintrunc (numeral k - 1) (bin_rest x) BIT bin_last x" + by (subst expand_Suc, rule bintrunc.simps) + +lemma sbintrunc_numeral: + "sbintrunc (numeral k) x = + sbintrunc (numeral k - 1) (bin_rest x) BIT bin_last x" + by (subst expand_Suc, rule sbintrunc.simps) -lemmas sbintrunc_pred_simps [simp] = - sbintrunc_minus_simps [of "number_of bin", simplified nobm1] for bin +lemma bintrunc_numeral_simps [simp]: + "bintrunc (numeral k) (numeral (Num.Bit0 w)) = + bintrunc (numeral k - 1) (numeral w) BIT 0" + "bintrunc (numeral k) (numeral (Num.Bit1 w)) = + bintrunc (numeral k - 1) (numeral w) BIT 1" + "bintrunc (numeral k) (neg_numeral (Num.Bit0 w)) = + bintrunc (numeral k - 1) (neg_numeral w) BIT 0" + "bintrunc (numeral k) (neg_numeral (Num.Bit1 w)) = + bintrunc (numeral k - 1) (neg_numeral (w + Num.One)) BIT 1" + "bintrunc (numeral k) 1 = 1" + by (simp_all add: bintrunc_numeral) -lemma no_bintr_alt: - "number_of (bintrunc n w) = w mod 2 ^ n" - by (simp add: number_of_eq bintrunc_mod2p) +lemma sbintrunc_numeral_simps [simp]: + "sbintrunc (numeral k) (numeral (Num.Bit0 w)) = + sbintrunc (numeral k - 1) (numeral w) BIT 0" + "sbintrunc (numeral k) (numeral (Num.Bit1 w)) = + sbintrunc (numeral k - 1) (numeral w) BIT 1" + "sbintrunc (numeral k) (neg_numeral (Num.Bit0 w)) = + sbintrunc (numeral k - 1) (neg_numeral w) BIT 0" + "sbintrunc (numeral k) (neg_numeral (Num.Bit1 w)) = + sbintrunc (numeral k - 1) (neg_numeral (w + Num.One)) BIT 1" + "sbintrunc (numeral k) 1 = 1" + by (simp_all add: sbintrunc_numeral) lemma no_bintr_alt1: "bintrunc n = (%w. w mod 2 ^ n :: int)" by (rule ext) (rule bintrunc_mod2p) @@ -620,19 +585,10 @@ apply (auto intro: int_mod_lem [THEN iffD1, symmetric]) done -lemma no_bintr: - "number_of (bintrunc n w) = (number_of w mod 2 ^ n :: int)" - by (simp add : bintrunc_mod2p number_of_eq) - lemma no_sbintr_alt2: "sbintrunc n = (%w. (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n :: int)" by (rule ext) (simp add : sbintrunc_mod2p) -lemma no_sbintr: - "number_of (sbintrunc n w) = - ((number_of w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n :: int)" - by (simp add : no_sbintr_alt2 number_of_eq) - lemma range_sbintrunc: "range (sbintrunc n) = {i. - (2 ^ n) <= i & i < 2 ^ n}" apply (unfold no_sbintr_alt2) @@ -692,21 +648,20 @@ lemmas m2pths = pos_mod_sign pos_mod_bound [OF zless2p] -lemma bintr_ge0: "(0 :: int) <= number_of (bintrunc n w)" - by (simp add : no_bintr m2pths) +lemma bintr_ge0: "0 \ bintrunc n w" + by (simp add: bintrunc_mod2p) -lemma bintr_lt2p: "number_of (bintrunc n w) < (2 ^ n :: int)" - by (simp add : no_bintr m2pths) +lemma bintr_lt2p: "bintrunc n w < 2 ^ n" + by (simp add: bintrunc_mod2p) -lemma bintr_Min: - "number_of (bintrunc n Int.Min) = (2 ^ n :: int) - 1" - by (simp add : no_bintr m1mod2k) +lemma bintr_Min: "bintrunc n -1 = 2 ^ n - 1" + by (simp add: bintrunc_mod2p m1mod2k) -lemma sbintr_ge: "(- (2 ^ n) :: int) <= number_of (sbintrunc n w)" - by (simp add : no_sbintr m2pths) +lemma sbintr_ge: "- (2 ^ n) \ sbintrunc n w" + by (simp add: sbintrunc_mod2p) -lemma sbintr_lt: "number_of (sbintrunc n w) < (2 ^ n :: int)" - by (simp add : no_sbintr m2pths) +lemma sbintr_lt: "sbintrunc n w < 2 ^ n" + by (simp add: sbintrunc_mod2p) lemma sign_Pls_ge_0: "(bin_sign bin = 0) = (bin >= (0 :: int))" @@ -716,8 +671,6 @@ "(bin_sign bin = -1) = (bin < (0 :: int))" unfolding bin_sign_def by simp -lemmas sign_Min_neg = trans [OF sign_Min_lt_0 neg_def [symmetric]] - lemma bin_rest_trunc: "(bin_rest (bintrunc n bin)) = bintrunc (n - 1) (bin_rest bin)" by (induct n arbitrary: bin) auto @@ -789,7 +742,7 @@ lemma [code]: "bin_split (Suc n) w = (let (w1, w2) = bin_split n (bin_rest w) in (w1, w2 BIT bin_last w))" "bin_split 0 w = (w, 0)" - by (simp_all add: Pls_def) + by simp_all primrec bin_cat :: "int \ nat \ int \ int" where Z: "bin_cat w 0 v = w" @@ -801,24 +754,17 @@ "0 < n \ f ^^ n = f \ f ^^ (n - 1)" by (cases n) simp_all -lemmas funpow_pred_simp [simp] = - funpow_minus_simp [of "number_of bin", simplified nobm1] for bin +lemma funpow_numeral [simp]: + "f ^^ numeral k = f \ f ^^ (numeral k - 1)" + by (subst expand_Suc, rule funpow.simps) -lemmas replicate_minus_simp = - trans [OF gen_minus [where f = "%n. replicate n x"] replicate.replicate_Suc] for x - -lemmas replicate_pred_simp [simp] = - replicate_minus_simp [of "number_of bin", simplified nobm1] for bin - -lemmas power_Suc_no [simp] = power_Suc [of "number_of a"] for a +lemma replicate_numeral [simp]: (* TODO: move to List.thy *) + "replicate (numeral k) x = x # replicate (numeral k - 1) x" + by (subst expand_Suc, rule replicate_Suc) lemmas power_minus_simp = trans [OF gen_minus [where f = "power f"] power_Suc] for f -lemmas power_pred_simp = - power_minus_simp [of "number_of bin", simplified nobm1] for bin -lemmas power_pred_simp_no [simp] = power_pred_simp [where f= "number_of f"] for f - lemma list_exhaust_size_gt0: assumes y: "\a list. y = a # list \ P" shows "0 < length y \ P" @@ -839,11 +785,6 @@ "y = xa # list ==> size y = Suc k ==> size list = k" by auto -lemma size_Cons_lem_eq_bin: - "y = xa # list ==> size y = number_of (Int.succ k) ==> - size list = number_of k" - by (auto simp: pred_def succ_def split add : split_if_asm) - lemmas ls_splits = prod.split prod.split_asm split_if_asm lemma not_B1_is_B0: "y \ (1::bit) \ y = (0::bit)" diff -r 29e92b644d6c -r f067afe98049 src/HOL/Word/Bool_List_Representation.thy --- a/src/HOL/Word/Bool_List_Representation.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Word/Bool_List_Representation.thy Mon Mar 26 10:56:56 2012 +0200 @@ -106,13 +106,13 @@ by (cases n) auto lemma bin_to_bl_aux_Bit0_minus_simp [simp]: - "0 < n ==> bin_to_bl_aux n (number_of (Int.Bit0 w)) bl = - bin_to_bl_aux (n - 1) (number_of w) (False # bl)" + "0 < n ==> bin_to_bl_aux n (numeral (Num.Bit0 w)) bl = + bin_to_bl_aux (n - 1) (numeral w) (False # bl)" by (cases n) auto lemma bin_to_bl_aux_Bit1_minus_simp [simp]: - "0 < n ==> bin_to_bl_aux n (number_of (Int.Bit1 w)) bl = - bin_to_bl_aux (n - 1) (number_of w) (True # bl)" + "0 < n ==> bin_to_bl_aux n (numeral (Num.Bit1 w)) bl = + bin_to_bl_aux (n - 1) (numeral w) (True # bl)" by (cases n) auto text {* Link between bin and bool list. *} @@ -632,8 +632,13 @@ lemmas takefill_minus_simps = takefill_Suc_cases [THEN [2] takefill_minus [symmetric, THEN trans]] -lemmas takefill_pred_simps [simp] = - takefill_minus_simps [where n="number_of bin", simplified nobm1] for bin +lemma takefill_numeral_Nil [simp]: + "takefill fill (numeral k) [] = fill # takefill fill (numeral k - 1) []" + by (subst expand_Suc, rule takefill_Suc_Nil) + +lemma takefill_numeral_Cons [simp]: + "takefill fill (numeral k) (x # xs) = x # takefill fill (numeral k - 1) xs" + by (subst expand_Suc, rule takefill_Suc_Cons) (* links with function bl_to_bin *) @@ -1031,11 +1036,11 @@ bin_split.Suc [THEN [2] bin_split_minus [symmetric, THEN trans]] lemma bin_split_pred_simp [simp]: - "(0::nat) < number_of bin \ - bin_split (number_of bin) w = - (let (w1, w2) = bin_split (number_of (Int.pred bin)) (bin_rest w) + "(0::nat) < numeral bin \ + bin_split (numeral bin) w = + (let (w1, w2) = bin_split (numeral bin - 1) (bin_rest w) in (w1, w2 BIT bin_last w))" - by (simp only: nobm1 bin_split_minus_simp) + by (simp only: bin_split_minus_simp) lemma bin_rsplit_aux_simp_alt: "bin_rsplit_aux n m c bs = diff -r 29e92b644d6c -r f067afe98049 src/HOL/Word/Misc_Numeric.thy --- a/src/HOL/Word/Misc_Numeric.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Word/Misc_Numeric.thy Mon Mar 26 10:56:56 2012 +0200 @@ -5,7 +5,7 @@ header {* Useful Numerical Lemmas *} theory Misc_Numeric -imports Main Parity +imports "~~/src/HOL/Main" "~~/src/HOL/Parity" begin lemma the_elemI: "y = {x} ==> the_elem y = x" @@ -31,13 +31,6 @@ lemma sum_imp_diff: "j = k + i ==> j - i = (k :: nat)" by arith -lemma nobm1: - "0 < (number_of w :: nat) ==> - number_of w - (1 :: nat) = number_of (Int.pred w)" - apply (unfold nat_number_of_def One_nat_def nat_1 [symmetric] pred_def) - apply (simp add: number_of_eq nat_diff_distrib [symmetric]) - done - lemma zless2: "0 < (2 :: int)" by arith lemmas zless2p [simp] = zless2 [THEN zero_less_power] @@ -46,7 +39,6 @@ lemmas pos_mod_sign2 = zless2 [THEN pos_mod_sign [where b = "2::int"]] lemmas pos_mod_bound2 = zless2 [THEN pos_mod_bound [where b = "2::int"]] --- "the inverse(s) of @{text number_of}" lemma nmod2: "n mod (2::int) = 0 | n mod 2 = 1" by arith lemma emep1: @@ -283,15 +275,6 @@ lemmas min_minus' [simp] = trans [OF min_max.inf_commute min_minus] -lemma nat_no_eq_iff: - "(number_of b :: int) >= 0 ==> (number_of c :: int) >= 0 ==> - (number_of b = (number_of c :: nat)) = (b = c)" - apply (unfold nat_number_of_def) - apply safe - apply (drule (2) eq_nat_nat_iff [THEN iffD1]) - apply (simp add: number_of_eq) - done - lemmas dme = box_equals [OF div_mod_equality add_0_right add_0_right] lemmas dtle = xtr3 [OF dme [symmetric] le_add1] lemmas th2 = order_trans [OF order_refl [THEN [2] mult_le_mono] dtle] diff -r 29e92b644d6c -r f067afe98049 src/HOL/Word/Word.thy --- a/src/HOL/Word/Word.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/Word/Word.thy Mon Mar 26 10:56:56 2012 +0200 @@ -20,17 +20,64 @@ typedef (open) 'a word = "{(0::int) ..< 2^len_of TYPE('a::len0)}" morphisms uint Abs_word by auto +lemma uint_nonnegative: + "0 \ uint w" + using word.uint [of w] by simp + +lemma uint_bounded: + fixes w :: "'a::len0 word" + shows "uint w < 2 ^ len_of TYPE('a)" + using word.uint [of w] by simp + +lemma uint_idem: + fixes w :: "'a::len0 word" + shows "uint w mod 2 ^ len_of TYPE('a) = uint w" + using uint_nonnegative uint_bounded by (rule mod_pos_pos_trivial) + definition word_of_int :: "int \ 'a\len0 word" where -- {* representation of words using unsigned or signed bins, only difference in these is the type class *} - "word_of_int w = Abs_word (bintrunc (len_of TYPE ('a)) w)" - -lemma uint_word_of_int [code]: "uint (word_of_int w \ 'a\len0 word) = w mod 2 ^ len_of TYPE('a)" - by (auto simp add: word_of_int_def bintrunc_mod2p intro: Abs_word_inverse) - -code_datatype word_of_int - -subsection {* Random instance *} + "word_of_int k = Abs_word (k mod 2 ^ len_of TYPE('a))" + +lemma uint_word_of_int: + "uint (word_of_int k :: 'a::len0 word) = k mod 2 ^ len_of TYPE('a)" + by (auto simp add: word_of_int_def intro: Abs_word_inverse) + +lemma word_of_int_uint: + "word_of_int (uint w) = w" + by (simp add: word_of_int_def uint_idem uint_inverse) + +lemma word_uint_eq_iff: + "a = b \ uint a = uint b" + by (simp add: uint_inject) + +lemma word_uint_eqI: + "uint a = uint b \ a = b" + by (simp add: word_uint_eq_iff) + + +subsection {* Basic code generation setup *} + +definition Word :: "int \ 'a::len0 word" +where + [code_post]: "Word = word_of_int" + +lemma [code abstype]: + "Word (uint w) = w" + by (simp add: Word_def word_of_int_uint) + +declare uint_word_of_int [code abstract] + +instantiation word :: (len0) equal +begin + +definition equal_word :: "'a word \ 'a word \ bool" where + "equal_word k l \ HOL.equal (uint k) (uint l)" + +instance proof +qed (simp add: equal equal_word_def word_uint_eq_iff) + +end notation fcomp (infixl "\>" 60) notation scomp (infixl "\\" 60) @@ -39,7 +86,7 @@ begin definition - "random_word i = Random.range (max i (2 ^ len_of TYPE('a))) \\ (\k. Pair ( + "random_word i = Random.range i \\ (\k. Pair ( let j = word_of_int (Code_Numeral.int_of k) :: 'a word in (j, \_::unit. Code_Evaluation.term_of j)))" @@ -193,7 +240,7 @@ where "word_pred a = word_of_int (uint a - 1)" -instantiation word :: (len0) "{number, Divides.div, comm_monoid_mult, comm_ring}" +instantiation word :: (len0) "{neg_numeral, Divides.div, comm_monoid_mult, comm_ring}" begin definition @@ -220,9 +267,6 @@ definition word_mod_def: "a mod b = word_of_int (uint a mod uint b)" -definition - word_number_of_def: "number_of w = word_of_int w" - lemmas word_arith_wis = word_add_def word_sub_wi word_mult_def word_minus_def word_succ_def word_pred_def word_0_wi word_1_wi @@ -268,9 +312,6 @@ apply (simp add: word_of_nat wi_hom_sub) done -instance word :: (len) number_ring - by (default, simp add: word_number_of_def word_of_int) - definition udvd :: "'a::len word => 'a::len word => bool" (infixl "udvd" 50) where "a udvd b = (EX n>=0. uint b = n * uint a)" @@ -284,7 +325,7 @@ word_le_def: "a \ b \ uint a \ uint b" definition - word_less_def: "x < y \ x \ y \ x \ (y \ 'a word)" + word_less_def: "a < b \ uint a < uint b" instance by default (auto simp: word_less_def word_le_def) @@ -504,40 +545,55 @@ lemmas td_sint = word_sint.td -lemma word_number_of_alt: - "number_of b = word_of_int (number_of b)" - by (simp add: number_of_eq word_number_of_def) - -declare word_number_of_alt [symmetric, code_abbrev] - -lemma word_no_wi: "number_of = word_of_int" - by (auto simp: word_number_of_def) - lemma to_bl_def': "(to_bl :: 'a :: len0 word => bool list) = bin_to_bl (len_of TYPE('a)) o uint" by (auto simp: to_bl_def) -lemmas word_reverse_no_def [simp] = word_reverse_def [of "number_of w"] for w +lemmas word_reverse_no_def [simp] = word_reverse_def [of "numeral w"] for w lemma uints_mod: "uints n = range (\w. w mod 2 ^ n)" by (fact uints_def [unfolded no_bintr_alt1]) +lemma word_numeral_alt: + "numeral b = word_of_int (numeral b)" + by (induct b, simp_all only: numeral.simps word_of_int_homs) + +declare word_numeral_alt [symmetric, code_abbrev] + +lemma word_neg_numeral_alt: + "neg_numeral b = word_of_int (neg_numeral b)" + by (simp only: neg_numeral_def word_numeral_alt wi_hom_neg) + +declare word_neg_numeral_alt [symmetric, code_abbrev] + lemma uint_bintrunc [simp]: - "uint (number_of bin :: 'a word) = - bintrunc (len_of TYPE ('a :: len0)) (number_of bin)" - unfolding word_number_of_alt by (rule word_ubin.eq_norm) + "uint (numeral bin :: 'a word) = + bintrunc (len_of TYPE ('a :: len0)) (numeral bin)" + unfolding word_numeral_alt by (rule word_ubin.eq_norm) + +lemma uint_bintrunc_neg [simp]: "uint (neg_numeral bin :: 'a word) = + bintrunc (len_of TYPE ('a :: len0)) (neg_numeral bin)" + by (simp only: word_neg_numeral_alt word_ubin.eq_norm) lemma sint_sbintrunc [simp]: - "sint (number_of bin :: 'a word) = - sbintrunc (len_of TYPE ('a :: len) - 1) (number_of bin)" - unfolding word_number_of_alt by (rule word_sbin.eq_norm) + "sint (numeral bin :: 'a word) = + sbintrunc (len_of TYPE ('a :: len) - 1) (numeral bin)" + by (simp only: word_numeral_alt word_sbin.eq_norm) + +lemma sint_sbintrunc_neg [simp]: "sint (neg_numeral bin :: 'a word) = + sbintrunc (len_of TYPE ('a :: len) - 1) (neg_numeral bin)" + by (simp only: word_neg_numeral_alt word_sbin.eq_norm) lemma unat_bintrunc [simp]: - "unat (number_of bin :: 'a :: len0 word) = - nat (bintrunc (len_of TYPE('a)) (number_of bin))" - unfolding unat_def nat_number_of_def - by (simp only: uint_bintrunc) + "unat (numeral bin :: 'a :: len0 word) = + nat (bintrunc (len_of TYPE('a)) (numeral bin))" + by (simp only: unat_def uint_bintrunc) + +lemma unat_bintrunc_neg [simp]: + "unat (neg_numeral bin :: 'a :: len0 word) = + nat (bintrunc (len_of TYPE('a)) (neg_numeral bin))" + by (simp only: unat_def uint_bintrunc_neg) lemma size_0_eq: "size (w :: 'a :: len0 word) = 0 \ v = w" apply (unfold word_size) @@ -562,7 +618,7 @@ lemma sign_uint_Pls [simp]: "bin_sign (uint x) = 0" - by (simp add: sign_Pls_ge_0 number_of_eq) + by (simp add: sign_Pls_ge_0) lemma uint_m2p_neg: "uint (x::'a::len0 word) - 2 ^ len_of TYPE('a) < 0" by (simp only: diff_less_0_iff_less uint_lt2p) @@ -581,35 +637,43 @@ lemma uint_nat: "uint w = int (unat w)" unfolding unat_def by auto -lemma uint_number_of: - "uint (number_of b :: 'a :: len0 word) = number_of b mod 2 ^ len_of TYPE('a)" - unfolding word_number_of_alt +lemma uint_numeral: + "uint (numeral b :: 'a :: len0 word) = numeral b mod 2 ^ len_of TYPE('a)" + unfolding word_numeral_alt by (simp only: int_word_uint) -lemma unat_number_of: - "bin_sign (number_of b) = 0 \ - unat (number_of b::'a::len0 word) = number_of b mod 2 ^ len_of TYPE ('a)" +lemma uint_neg_numeral: + "uint (neg_numeral b :: 'a :: len0 word) = neg_numeral b mod 2 ^ len_of TYPE('a)" + unfolding word_neg_numeral_alt + by (simp only: int_word_uint) + +lemma unat_numeral: + "unat (numeral b::'a::len0 word) = numeral b mod 2 ^ len_of TYPE ('a)" apply (unfold unat_def) - apply (clarsimp simp only: uint_number_of) + apply (clarsimp simp only: uint_numeral) apply (rule nat_mod_distrib [THEN trans]) - apply (erule sign_Pls_ge_0 [THEN iffD1]) + apply (rule zero_le_numeral) apply (simp_all add: nat_power_eq) done -lemma sint_number_of: "sint (number_of b :: 'a :: len word) = (number_of b + +lemma sint_numeral: "sint (numeral b :: 'a :: len word) = (numeral b + 2 ^ (len_of TYPE('a) - 1)) mod 2 ^ len_of TYPE('a) - 2 ^ (len_of TYPE('a) - 1)" - unfolding word_number_of_alt by (rule int_word_sint) - -lemma word_of_int_0 [simp]: "word_of_int 0 = 0" + unfolding word_numeral_alt by (rule int_word_sint) + +lemma word_of_int_0 [simp, code_post]: "word_of_int 0 = 0" unfolding word_0_wi .. -lemma word_of_int_1 [simp]: "word_of_int 1 = 1" +lemma word_of_int_1 [simp, code_post]: "word_of_int 1 = 1" unfolding word_1_wi .. -lemma word_of_int_bin [simp] : - "(word_of_int (number_of bin) :: 'a :: len0 word) = (number_of bin)" - unfolding word_number_of_alt .. +lemma word_of_int_numeral [simp] : + "(word_of_int (numeral bin) :: 'a :: len0 word) = (numeral bin)" + unfolding word_numeral_alt .. + +lemma word_of_int_neg_numeral [simp]: + "(word_of_int (neg_numeral bin) :: 'a :: len0 word) = (neg_numeral bin)" + unfolding neg_numeral_def word_numeral_alt wi_hom_syms .. lemma word_int_case_wi: "word_int_case f (word_of_int i :: 'b word) = @@ -728,7 +792,7 @@ unfolding word_reverse_def by (simp add : word_bl.Abs_inverse) lemma word_rev_gal: "word_reverse w = u \ word_reverse u = w" - by auto + by (metis word_rev_rev) lemma word_rev_gal': "u = word_reverse w \ w = word_reverse u" by simp @@ -762,8 +826,8 @@ done lemma no_of_bl: - "(number_of bin ::'a::len0 word) = of_bl (bin_to_bl (len_of TYPE ('a)) bin)" - unfolding word_size of_bl_def by (simp add: word_number_of_def) + "(numeral bin ::'a::len0 word) = of_bl (bin_to_bl (len_of TYPE ('a)) (numeral bin))" + unfolding of_bl_def by simp lemma uint_bl: "to_bl w = bin_to_bl (size w) (uint w)" unfolding word_size to_bl_def by auto @@ -775,9 +839,15 @@ "to_bl (word_of_int bin::'a::len0 word) = bin_to_bl (len_of TYPE('a)) bin" unfolding uint_bl by (clarsimp simp add: word_ubin.eq_norm word_size) -lemma to_bl_no_bin [simp]: - "to_bl (number_of bin::'a::len0 word) = bin_to_bl (len_of TYPE('a)) (number_of bin)" - unfolding word_number_of_alt by (rule to_bl_of_bin) +lemma to_bl_numeral [simp]: + "to_bl (numeral bin::'a::len0 word) = + bin_to_bl (len_of TYPE('a)) (numeral bin)" + unfolding word_numeral_alt by (rule to_bl_of_bin) + +lemma to_bl_neg_numeral [simp]: + "to_bl (neg_numeral bin::'a::len0 word) = + bin_to_bl (len_of TYPE('a)) (neg_numeral bin)" + unfolding word_neg_numeral_alt by (rule to_bl_of_bin) lemma to_bl_to_bin [simp] : "bl_to_bin (to_bl w) = uint w" unfolding uint_bl by (simp add : word_size) @@ -803,35 +873,29 @@ by (auto simp add : uints_unats image_iff) lemmas bintr_num = word_ubin.norm_eq_iff - [of "number_of a" "number_of b", symmetric, folded word_number_of_alt] for a b + [of "numeral a" "numeral b", symmetric, folded word_numeral_alt] for a b lemmas sbintr_num = word_sbin.norm_eq_iff - [of "number_of a" "number_of b", symmetric, folded word_number_of_alt] for a b - -lemmas num_of_bintr = word_ubin.Abs_norm [folded word_number_of_def] -lemmas num_of_sbintr = word_sbin.Abs_norm [folded word_number_of_def] - -(* don't add these to simpset, since may want bintrunc n w to be simplified; - may want these in reverse, but loop as simp rules, so use following *) + [of "numeral a" "numeral b", symmetric, folded word_numeral_alt] for a b lemma num_of_bintr': - "bintrunc (len_of TYPE('a :: len0)) (number_of a) = (number_of b) \ - number_of a = (number_of b :: 'a word)" + "bintrunc (len_of TYPE('a :: len0)) (numeral a) = (numeral b) \ + numeral a = (numeral b :: 'a word)" unfolding bintr_num by (erule subst, simp) lemma num_of_sbintr': - "sbintrunc (len_of TYPE('a :: len) - 1) (number_of a) = (number_of b) \ - number_of a = (number_of b :: 'a word)" + "sbintrunc (len_of TYPE('a :: len) - 1) (numeral a) = (numeral b) \ + numeral a = (numeral b :: 'a word)" unfolding sbintr_num by (erule subst, simp) lemma num_abs_bintr: - "(number_of x :: 'a word) = - word_of_int (bintrunc (len_of TYPE('a::len0)) (number_of x))" - by (simp only: word_ubin.Abs_norm word_number_of_alt) + "(numeral x :: 'a word) = + word_of_int (bintrunc (len_of TYPE('a::len0)) (numeral x))" + by (simp only: word_ubin.Abs_norm word_numeral_alt) lemma num_abs_sbintr: - "(number_of x :: 'a word) = - word_of_int (sbintrunc (len_of TYPE('a::len) - 1) (number_of x))" - by (simp only: word_sbin.Abs_norm word_number_of_alt) + "(numeral x :: 'a word) = + word_of_int (sbintrunc (len_of TYPE('a::len) - 1) (numeral x))" + by (simp only: word_sbin.Abs_norm word_numeral_alt) (** cast - note, no arg for new length, as it's determined by type of result, thus in "cast w = w, the type means cast to length of w! **) @@ -856,13 +920,14 @@ (* for literal u(s)cast *) lemma ucast_bintr [simp]: - "ucast (number_of w ::'a::len0 word) = - word_of_int (bintrunc (len_of TYPE('a)) (number_of w))" + "ucast (numeral w ::'a::len0 word) = + word_of_int (bintrunc (len_of TYPE('a)) (numeral w))" unfolding ucast_def by simp +(* TODO: neg_numeral *) lemma scast_sbintr [simp]: - "scast (number_of w ::'a::len word) = - word_of_int (sbintrunc (len_of TYPE('a) - Suc 0) (number_of w))" + "scast (numeral w ::'a::len word) = + word_of_int (sbintrunc (len_of TYPE('a) - Suc 0) (numeral w))" unfolding scast_def by simp lemma source_size: "source_size (c::'a::len0 word \ _) = len_of TYPE('a)" @@ -1016,8 +1081,8 @@ done lemma ucast_down_no [OF refl]: - "uc = ucast \ is_down uc \ uc (number_of bin) = number_of bin" - unfolding word_number_of_alt by clarify (rule ucast_down_wi) + "uc = ucast \ is_down uc \ uc (numeral bin) = numeral bin" + unfolding word_numeral_alt by clarify (rule ucast_down_wi) lemma ucast_down_bl [OF refl]: "uc = ucast \ is_down uc \ uc (of_bl bl) = of_bl bl" @@ -1028,19 +1093,6 @@ lemmas word_log_defs = word_and_def word_or_def word_xor_def word_not_def -text {* Executable equality *} - -instantiation word :: (len0) equal -begin - -definition equal_word :: "'a word \ 'a word \ bool" where - "equal_word k l \ HOL.equal (uint k) (uint l)" - -instance proof -qed (simp add: equal equal_word_def) - -end - subsection {* Word Arithmetic *} @@ -1057,33 +1109,23 @@ "0 \ n \ uint b = n * uint a \ a udvd b" by (auto simp: udvd_def) -lemmas word_div_no [simp] = word_div_def [of "number_of a" "number_of b"] for a b - -lemmas word_mod_no [simp] = word_mod_def [of "number_of a" "number_of b"] for a b - -lemmas word_less_no [simp] = word_less_def [of "number_of a" "number_of b"] for a b - -lemmas word_le_no [simp] = word_le_def [of "number_of a" "number_of b"] for a b - -lemmas word_sless_no [simp] = word_sless_def [of "number_of a" "number_of b"] for a b - -lemmas word_sle_no [simp] = word_sle_def [of "number_of a" "number_of b"] for a b - -(* following two are available in class number_ring, - but convenient to have them here here; - note - the number_ring versions, numeral_0_eq_0 and numeral_1_eq_1 - are in the default simpset, so to use the automatic simplifications for - (eg) sint (number_of bin) on sint 1, must do - (simp add: word_1_no del: numeral_1_eq_1) - *) -lemma word_0_no: "(0::'a::len0 word) = Numeral0" - by (simp add: word_number_of_alt) +lemmas word_div_no [simp] = word_div_def [of "numeral a" "numeral b"] for a b + +lemmas word_mod_no [simp] = word_mod_def [of "numeral a" "numeral b"] for a b + +lemmas word_less_no [simp] = word_less_def [of "numeral a" "numeral b"] for a b + +lemmas word_le_no [simp] = word_le_def [of "numeral a" "numeral b"] for a b + +lemmas word_sless_no [simp] = word_sless_def [of "numeral a" "numeral b"] for a b + +lemmas word_sle_no [simp] = word_sle_def [of "numeral a" "numeral b"] for a b lemma word_1_no: "(1::'a::len0 word) = Numeral1" - by (simp add: word_number_of_alt) + by (simp add: word_numeral_alt) lemma word_m1_wi: "-1 = word_of_int -1" - by (rule word_number_of_alt) + by (rule word_neg_numeral_alt) lemma word_0_bl [simp]: "of_bl [] = 0" unfolding of_bl_def by simp @@ -1195,17 +1237,18 @@ lemmas uint_mod_alt = word_mod_def [THEN trans [OF uint_cong int_word_uint]] lemma word_pred_0_n1: "word_pred 0 = word_of_int -1" - unfolding word_pred_def uint_eq_0 pred_def by simp + unfolding word_pred_def uint_eq_0 by simp lemma succ_pred_no [simp]: - "word_succ (number_of bin) = number_of (Int.succ bin) & - word_pred (number_of bin) = number_of (Int.pred bin)" - unfolding word_number_of_def Int.succ_def Int.pred_def - by (simp add: word_of_int_homs) + "word_succ (numeral w) = numeral w + 1" + "word_pred (numeral w) = numeral w - 1" + "word_succ (neg_numeral w) = neg_numeral w + 1" + "word_pred (neg_numeral w) = neg_numeral w - 1" + unfolding word_succ_p1 word_pred_m1 by simp_all lemma word_sp_01 [simp] : "word_succ -1 = 0 & word_succ 0 = 1 & word_pred 0 = -1 & word_pred 1 = 0" - unfolding word_0_no word_1_no by simp + unfolding word_succ_p1 word_pred_m1 by simp_all (* alternative approach to lifting arithmetic equalities *) lemma word_of_int_Ex: @@ -1230,10 +1273,10 @@ lemmas word_not_simps [simp] = word_zero_le [THEN leD] word_m1_ge [THEN leD] word_n1_ge [THEN leD] -lemma word_gt_0: "0 < y = (0 ~= (y :: 'a :: len0 word))" - unfolding word_less_def by auto - -lemmas word_gt_0_no [simp] = word_gt_0 [of "number_of y"] for y +lemma word_gt_0: "0 < y \ 0 \ (y :: 'a :: len0 word)" + by (simp add: less_le) + +lemmas word_gt_0_no [simp] = word_gt_0 [of "numeral y"] for y lemma word_sless_alt: "(a = 1, ie 'a :: len word *) lemma iszero_word_no [simp]: - "iszero (number_of bin :: 'a :: len word) = - iszero (bintrunc (len_of TYPE('a)) (number_of bin))" - using word_ubin.norm_eq_iff [where 'a='a, of "number_of bin" 0] + "iszero (numeral bin :: 'a :: len word) = + iszero (bintrunc (len_of TYPE('a)) (numeral bin))" + using word_ubin.norm_eq_iff [where 'a='a, of "numeral bin" 0] by (simp add: iszero_def [symmetric]) + +text {* Use @{text iszero} to simplify equalities between word numerals. *} + +lemmas word_eq_numeral_iff_iszero [simp] = + eq_numeral_iff_iszero [where 'a="'a::len word"] subsection "Word and nat" @@ -2023,10 +2071,10 @@ lemma of_bl_length_less: "length x = k \ k < len_of TYPE('a) \ (of_bl x :: 'a :: len word) < 2 ^ k" - apply (unfold of_bl_def word_less_alt word_number_of_alt) + apply (unfold of_bl_def word_less_alt word_numeral_alt) apply safe apply (simp (no_asm) add: word_of_int_power_hom word_uint.eq_norm - del: word_of_int_bin) + del: word_of_int_numeral) apply (simp add: mod_pos_pos_trivial) apply (subst mod_pos_pos_trivial) apply (rule bl_to_bin_ge0) @@ -2073,22 +2121,38 @@ unfolding word_log_defs wils1 by simp_all lemma word_no_log_defs [simp]: - "NOT number_of a = (number_of (NOT a) :: 'a::len0 word)" - "number_of a AND number_of b = (number_of (a AND b) :: 'a word)" - "number_of a OR number_of b = (number_of (a OR b) :: 'a word)" - "number_of a XOR number_of b = (number_of (a XOR b) :: 'a word)" - unfolding word_no_wi word_wi_log_defs by simp_all + "NOT (numeral a) = word_of_int (NOT (numeral a))" + "NOT (neg_numeral a) = word_of_int (NOT (neg_numeral a))" + "numeral a AND numeral b = word_of_int (numeral a AND numeral b)" + "numeral a AND neg_numeral b = word_of_int (numeral a AND neg_numeral b)" + "neg_numeral a AND numeral b = word_of_int (neg_numeral a AND numeral b)" + "neg_numeral a AND neg_numeral b = word_of_int (neg_numeral a AND neg_numeral b)" + "numeral a OR numeral b = word_of_int (numeral a OR numeral b)" + "numeral a OR neg_numeral b = word_of_int (numeral a OR neg_numeral b)" + "neg_numeral a OR numeral b = word_of_int (neg_numeral a OR numeral b)" + "neg_numeral a OR neg_numeral b = word_of_int (neg_numeral a OR neg_numeral b)" + "numeral a XOR numeral b = word_of_int (numeral a XOR numeral b)" + "numeral a XOR neg_numeral b = word_of_int (numeral a XOR neg_numeral b)" + "neg_numeral a XOR numeral b = word_of_int (neg_numeral a XOR numeral b)" + "neg_numeral a XOR neg_numeral b = word_of_int (neg_numeral a XOR neg_numeral b)" + unfolding word_numeral_alt word_neg_numeral_alt word_wi_log_defs by simp_all text {* Special cases for when one of the arguments equals 1. *} lemma word_bitwise_1_simps [simp]: "NOT (1::'a::len0 word) = -2" - "(1::'a word) AND number_of b = number_of (Int.Bit1 Int.Pls AND b)" - "number_of a AND (1::'a word) = number_of (a AND Int.Bit1 Int.Pls)" - "(1::'a word) OR number_of b = number_of (Int.Bit1 Int.Pls OR b)" - "number_of a OR (1::'a word) = number_of (a OR Int.Bit1 Int.Pls)" - "(1::'a word) XOR number_of b = number_of (Int.Bit1 Int.Pls XOR b)" - "number_of a XOR (1::'a word) = number_of (a XOR Int.Bit1 Int.Pls)" + "1 AND numeral b = word_of_int (1 AND numeral b)" + "1 AND neg_numeral b = word_of_int (1 AND neg_numeral b)" + "numeral a AND 1 = word_of_int (numeral a AND 1)" + "neg_numeral a AND 1 = word_of_int (neg_numeral a AND 1)" + "1 OR numeral b = word_of_int (1 OR numeral b)" + "1 OR neg_numeral b = word_of_int (1 OR neg_numeral b)" + "numeral a OR 1 = word_of_int (numeral a OR 1)" + "neg_numeral a OR 1 = word_of_int (neg_numeral a OR 1)" + "1 XOR numeral b = word_of_int (1 XOR numeral b)" + "1 XOR neg_numeral b = word_of_int (1 XOR neg_numeral b)" + "numeral a XOR 1 = word_of_int (numeral a XOR 1)" + "neg_numeral a XOR 1 = word_of_int (neg_numeral a XOR 1)" unfolding word_1_no word_no_log_defs by simp_all lemma uint_or: "uint (x OR y) = (uint x) OR (uint y)" @@ -2123,10 +2187,15 @@ unfolding word_test_bit_def by (simp add: nth_bintr [symmetric] word_ubin.eq_norm) -lemma test_bit_no [simp]: - "(number_of w :: 'a::len0 word) !! n \ - n < len_of TYPE('a) \ bin_nth (number_of w) n" - unfolding word_number_of_alt test_bit_wi .. +lemma test_bit_numeral [simp]: + "(numeral w :: 'a::len0 word) !! n \ + n < len_of TYPE('a) \ bin_nth (numeral w) n" + unfolding word_numeral_alt test_bit_wi .. + +lemma test_bit_neg_numeral [simp]: + "(neg_numeral w :: 'a::len0 word) !! n \ + n < len_of TYPE('a) \ bin_nth (neg_numeral w) n" + unfolding word_neg_numeral_alt test_bit_wi .. lemma test_bit_1 [simp]: "(1::'a::len word) !! n \ n = 0" unfolding word_1_wi test_bit_wi by auto @@ -2134,6 +2203,9 @@ lemma nth_0 [simp]: "~ (0::'a::len0 word) !! n" unfolding word_test_bit_def by simp +lemma nth_minus1 [simp]: "(-1::'a::len0 word) !! n \ n < len_of TYPE('a)" + unfolding word_test_bit_def by (simp add: nth_bintr) + (* get from commutativity, associativity etc of int_and etc to same for word_and etc *) @@ -2294,9 +2366,13 @@ "msb (word_of_int x::'a::len word) = bin_nth x (len_of TYPE('a) - 1)" unfolding word_msb_def by (simp add: word_sbin.eq_norm bin_sign_lem) -lemma word_msb_no [simp]: - "msb (number_of w::'a::len word) = bin_nth (number_of w) (len_of TYPE('a) - 1)" - unfolding word_number_of_alt by (rule msb_word_of_int) +lemma word_msb_numeral [simp]: + "msb (numeral w::'a::len word) = bin_nth (numeral w) (len_of TYPE('a) - 1)" + unfolding word_numeral_alt by (rule msb_word_of_int) + +lemma word_msb_neg_numeral [simp]: + "msb (neg_numeral w::'a::len word) = bin_nth (neg_numeral w) (len_of TYPE('a) - 1)" + unfolding word_neg_numeral_alt by (rule msb_word_of_int) lemma word_msb_0 [simp]: "\ msb (0::'a::len word)" unfolding word_msb_def by simp @@ -2420,9 +2496,13 @@ unfolding sint_uint l_def by (clarsimp simp add: nth_sbintr word_test_bit_def [symmetric]) -lemma word_lsb_no [simp]: - "lsb (number_of bin :: 'a :: len word) = (bin_last (number_of bin) = 1)" - unfolding word_lsb_alt test_bit_no by auto +lemma word_lsb_numeral [simp]: + "lsb (numeral bin :: 'a :: len word) = (bin_last (numeral bin) = 1)" + unfolding word_lsb_alt test_bit_numeral by simp + +lemma word_lsb_neg_numeral [simp]: + "lsb (neg_numeral bin :: 'a :: len word) = (bin_last (neg_numeral bin) = 1)" + unfolding word_lsb_alt test_bit_neg_numeral by simp lemma set_bit_word_of_int: "set_bit (word_of_int x) n b = word_of_int (bin_sc n (if b then 1 else 0) x)" @@ -2431,10 +2511,15 @@ apply (simp add: word_size bin_nth_sc_gen word_ubin.eq_norm nth_bintr) done -lemma word_set_no [simp]: - "set_bit (number_of bin::'a::len0 word) n b = - word_of_int (bin_sc n (if b then 1 else 0) (number_of bin))" - unfolding word_number_of_alt by (rule set_bit_word_of_int) +lemma word_set_numeral [simp]: + "set_bit (numeral bin::'a::len0 word) n b = + word_of_int (bin_sc n (if b then 1 else 0) (numeral bin))" + unfolding word_numeral_alt by (rule set_bit_word_of_int) + +lemma word_set_neg_numeral [simp]: + "set_bit (neg_numeral bin::'a::len0 word) n b = + word_of_int (bin_sc n (if b then 1 else 0) (neg_numeral bin))" + unfolding word_neg_numeral_alt by (rule set_bit_word_of_int) lemma word_set_bit_0 [simp]: "set_bit 0 n b = word_of_int (bin_sc n (if b then 1 else 0) 0)" @@ -2445,11 +2530,11 @@ unfolding word_1_wi by (rule set_bit_word_of_int) lemma setBit_no [simp]: - "setBit (number_of bin) n = word_of_int (bin_sc n 1 (number_of bin))" + "setBit (numeral bin) n = word_of_int (bin_sc n 1 (numeral bin))" by (simp add: setBit_def) lemma clearBit_no [simp]: - "clearBit (number_of bin) n = word_of_int (bin_sc n 0 (number_of bin))" + "clearBit (numeral bin) n = word_of_int (bin_sc n 0 (numeral bin))" by (simp add: clearBit_def) lemma to_bl_n1: @@ -2512,7 +2597,7 @@ apply (rule word_ubin.norm_eq_iff [THEN iffD1]) apply (rule box_equals) apply (rule_tac [2] bintr_ariths (1))+ - apply (clarsimp simp add : number_of_is_id) + apply simp apply simp done @@ -2547,15 +2632,19 @@ lemma shiftl1_wi [simp]: "shiftl1 (word_of_int w) = word_of_int (w BIT 0)" unfolding shiftl1_def - apply (simp only: word_ubin.norm_eq_iff [symmetric] word_ubin.eq_norm) + apply (simp add: word_ubin.norm_eq_iff [symmetric] word_ubin.eq_norm) apply (subst refl [THEN bintrunc_BIT_I, symmetric]) apply (subst bintrunc_bintrunc_min) apply simp done -lemma shiftl1_number [simp] : - "shiftl1 (number_of w) = number_of (Int.Bit0 w)" - unfolding word_number_of_alt shiftl1_wi by simp +lemma shiftl1_numeral [simp]: + "shiftl1 (numeral w) = numeral (Num.Bit0 w)" + unfolding word_numeral_alt shiftl1_wi by simp + +lemma shiftl1_neg_numeral [simp]: + "shiftl1 (neg_numeral w) = neg_numeral (Num.Bit0 w)" + unfolding word_neg_numeral_alt shiftl1_wi by simp lemma shiftl1_0 [simp] : "shiftl1 0 = 0" unfolding shiftl1_def by simp @@ -2704,8 +2793,8 @@ subsubsection "shift functions in terms of lists of bools" -lemmas bshiftr1_no_bin [simp] = - bshiftr1_def [where w="number_of w", unfolded to_bl_no_bin] for w +lemmas bshiftr1_numeral [simp] = + bshiftr1_def [where w="numeral w", unfolded to_bl_numeral] for w lemma bshiftr1_bl: "to_bl (bshiftr1 b w) = b # butlast (to_bl w)" unfolding bshiftr1_def by (rule word_bl.Abs_inverse) simp @@ -2858,7 +2947,7 @@ finally show ?thesis . qed -lemmas shiftl_number [simp] = shiftl_def [where w="number_of w"] for w +lemmas shiftl_numeral [simp] = shiftl_def [where w="numeral w"] for w lemma bl_shiftl: "to_bl (w << n) = drop n (to_bl w) @ replicate (min (size w) n) False" @@ -2885,27 +2974,29 @@ by (induct n) (auto simp: shiftl1_2t) lemma shiftr1_bintr [simp]: - "(shiftr1 (number_of w) :: 'a :: len0 word) = - word_of_int (bin_rest (bintrunc (len_of TYPE ('a)) (number_of w)))" - unfolding shiftr1_def word_number_of_alt + "(shiftr1 (numeral w) :: 'a :: len0 word) = + word_of_int (bin_rest (bintrunc (len_of TYPE ('a)) (numeral w)))" + unfolding shiftr1_def word_numeral_alt by (simp add: word_ubin.eq_norm) lemma sshiftr1_sbintr [simp]: - "(sshiftr1 (number_of w) :: 'a :: len word) = - word_of_int (bin_rest (sbintrunc (len_of TYPE ('a) - 1) (number_of w)))" - unfolding sshiftr1_def word_number_of_alt + "(sshiftr1 (numeral w) :: 'a :: len word) = + word_of_int (bin_rest (sbintrunc (len_of TYPE ('a) - 1) (numeral w)))" + unfolding sshiftr1_def word_numeral_alt by (simp add: word_sbin.eq_norm) lemma shiftr_no [simp]: - "(number_of w::'a::len0 word) >> n = word_of_int - ((bin_rest ^^ n) (bintrunc (len_of TYPE('a)) (number_of w)))" + (* FIXME: neg_numeral *) + "(numeral w::'a::len0 word) >> n = word_of_int + ((bin_rest ^^ n) (bintrunc (len_of TYPE('a)) (numeral w)))" apply (rule word_eqI) apply (auto simp: nth_shiftr nth_rest_power_bin nth_bintr word_size) done lemma sshiftr_no [simp]: - "(number_of w::'a::len word) >>> n = word_of_int - ((bin_rest ^^ n) (sbintrunc (len_of TYPE('a) - 1) (number_of w)))" + (* FIXME: neg_numeral *) + "(numeral w::'a::len word) >>> n = word_of_int + ((bin_rest ^^ n) (sbintrunc (len_of TYPE('a) - 1) (numeral w)))" apply (rule word_eqI) apply (auto simp: nth_sshiftr nth_rest_power_bin nth_sbintr word_size) apply (subgoal_tac "na + n = len_of TYPE('a) - Suc 0", simp, simp)+ @@ -3016,8 +3107,8 @@ lemma and_mask_wi: "word_of_int i AND mask n = word_of_int (bintrunc n i)" by (auto simp add: nth_bintr word_size word_ops_nth_size word_eq_iff) -lemma and_mask_no: "number_of i AND mask n = word_of_int (bintrunc n (number_of i))" - unfolding word_number_of_alt by (rule and_mask_wi) +lemma and_mask_no: "numeral i AND mask n = word_of_int (bintrunc n (numeral i))" + unfolding word_numeral_alt by (rule and_mask_wi) lemma bl_and_mask': "to_bl (w AND mask n :: 'a :: len word) = @@ -3046,7 +3137,7 @@ by (simp add: int_mod_lem eq_sym_conv) lemma mask_eq_iff: "(w AND mask n) = w <-> uint w < 2 ^ n" - apply (simp add: and_mask_bintr word_number_of_def) + apply (simp add: and_mask_bintr) apply (simp add: word_ubin.inverse_norm) apply (simp add: eq_mod_iff bintrunc_mod2p min_def) apply (fast intro!: lt2p_lem) @@ -3073,17 +3164,17 @@ lemma word_2p_lem: "n < size w \ w < 2 ^ n = (uint (w :: 'a :: len word) < 2 ^ n)" - apply (unfold word_size word_less_alt word_number_of_alt) + apply (unfold word_size word_less_alt word_numeral_alt) apply (clarsimp simp add: word_of_int_power_hom word_uint.eq_norm int_mod_eq' - simp del: word_of_int_bin) + simp del: word_of_int_numeral) done lemma less_mask_eq: "x < 2 ^ n \ x AND mask n = (x :: 'a :: len word)" - apply (unfold word_less_alt word_number_of_alt) + apply (unfold word_less_alt word_numeral_alt) apply (clarsimp simp add: and_mask_mod_2p word_of_int_power_hom word_uint.eq_norm - simp del: word_of_int_bin) + simp del: word_of_int_numeral) apply (drule xtr8 [rotated]) apply (rule int_mod_le) apply (auto simp add : mod_pos_pos_trivial) @@ -3126,7 +3217,7 @@ lemmas revcast_def' = revcast_def [simplified] lemmas revcast_def'' = revcast_def' [simplified word_size] -lemmas revcast_no_def [simp] = revcast_def' [where w="number_of w", unfolded word_size] for w +lemmas revcast_no_def [simp] = revcast_def' [where w="numeral w", unfolded word_size] for w lemma to_bl_revcast: "to_bl (revcast w :: 'a :: len0 word) = @@ -3240,13 +3331,13 @@ subsubsection "Slices" lemma slice1_no_bin [simp]: - "slice1 n (number_of w :: 'b word) = of_bl (takefill False n (bin_to_bl (len_of TYPE('b :: len0)) (number_of w)))" - by (simp add: slice1_def) + "slice1 n (numeral w :: 'b word) = of_bl (takefill False n (bin_to_bl (len_of TYPE('b :: len0)) (numeral w)))" + by (simp add: slice1_def) (* TODO: neg_numeral *) lemma slice_no_bin [simp]: - "slice n (number_of w :: 'b word) = of_bl (takefill False (len_of TYPE('b :: len0) - n) - (bin_to_bl (len_of TYPE('b :: len0)) (number_of w)))" - by (simp add: slice_def word_size) + "slice n (numeral w :: 'b word) = of_bl (takefill False (len_of TYPE('b :: len0) - n) + (bin_to_bl (len_of TYPE('b :: len0)) (numeral w)))" + by (simp add: slice_def word_size) (* TODO: neg_numeral *) lemma slice1_0 [simp] : "slice1 n 0 = 0" unfolding slice1_def by simp @@ -3383,9 +3474,9 @@ lemmas word_cat_bin' = word_cat_def lemma word_rsplit_no: - "(word_rsplit (number_of bin :: 'b :: len0 word) :: 'a word list) = + "(word_rsplit (numeral bin :: 'b :: len0 word) :: 'a word list) = map word_of_int (bin_rsplit (len_of TYPE('a :: len)) - (len_of TYPE('b), bintrunc (len_of TYPE('b)) (number_of bin)))" + (len_of TYPE('b), bintrunc (len_of TYPE('b)) (numeral bin)))" unfolding word_rsplit_def by (simp add: word_ubin.eq_norm) lemmas word_rsplit_no_cl [simp] = word_rsplit_no @@ -3580,15 +3671,14 @@ done lemmas word_cat_bl_no_bin [simp] = - word_cat_bl [where a="number_of a" - and b="number_of b", - unfolded to_bl_no_bin] - for a b + word_cat_bl [where a="numeral a" and b="numeral b", + unfolded to_bl_numeral] + for a b (* FIXME: negative numerals, 0 and 1 *) lemmas word_split_bl_no_bin [simp] = - word_split_bl_eq [where c="number_of c", unfolded to_bl_no_bin] for c - --- {* this odd result arises from the fact that the statement of the + word_split_bl_eq [where c="numeral c", unfolded to_bl_numeral] for c + +text {* this odd result arises from the fact that the statement of the result implies that the decoded words are of the same type, and therefore of the same length, as the original word *} @@ -3962,7 +4052,7 @@ lemma word_rotr_rev: "word_rotr n w = word_reverse (word_rotl n (word_reverse w))" - by (simp add: word_bl.Rep_inject [symmetric] to_bl_word_rev + by (simp only: word_bl.Rep_inject [symmetric] to_bl_word_rev to_bl_rotr to_bl_rotl rotater_rev) lemma word_roti_0 [simp]: "word_roti 0 w = w" @@ -4093,10 +4183,12 @@ unfolding word_roti_def by auto lemmas word_rotr_dt_no_bin' [simp] = - word_rotr_dt [where w="number_of w", unfolded to_bl_no_bin] for w + word_rotr_dt [where w="numeral w", unfolded to_bl_numeral] for w + (* FIXME: negative numerals, 0 and 1 *) lemmas word_rotl_dt_no_bin' [simp] = - word_rotl_dt [where w="number_of w", unfolded to_bl_no_bin] for w + word_rotl_dt [where w="numeral w", unfolded to_bl_numeral] for w + (* FIXME: negative numerals, 0 and 1 *) declare word_roti_def [simp] @@ -4119,8 +4211,7 @@ (simp add: max_word_def word_le_def int_word_uint int_mod_eq') lemma word_of_int_2p_len: "word_of_int (2 ^ len_of TYPE('a)) = (0::'a::len0 word)" - by (subst word_uint.Abs_norm [symmetric]) - (simp add: word_of_int_hom_syms) + by (subst word_uint.Abs_norm [symmetric]) simp lemma word_pow_0: "(2::'a::len word) ^ len_of TYPE('a) = 0" @@ -4304,10 +4395,7 @@ lemma word_neq_0_conv: fixes w :: "'a :: len word" shows "(w \ 0) = (0 < w)" -proof - - have "0 \ w" by (rule word_zero_le) - thus ?thesis by (auto simp add: word_less_def) -qed + unfolding word_gt_0 by simp lemma max_lt: "unat (max a b div c) = unat (max a b) div unat (c:: 'a :: len word)" @@ -4335,8 +4423,8 @@ "b <= a \ unat (a - b) = unat a - unat b" by (simp add: unat_def uint_sub_if_size word_le_def nat_diff_distrib) -lemmas word_less_sub1_numberof [simp] = word_less_sub1 [of "number_of w"] for w -lemmas word_le_sub1_numberof [simp] = word_le_sub1 [of "number_of w"] for w +lemmas word_less_sub1_numberof [simp] = word_less_sub1 [of "numeral w"] for w +lemmas word_le_sub1_numberof [simp] = word_le_sub1 [of "numeral w"] for w lemma word_of_int_minus: "word_of_int (2^len_of TYPE('a) - i) = (word_of_int (-i)::'a::len word)" @@ -4354,7 +4442,7 @@ lemma word_le_less_eq: "(x ::'z::len word) \ y = (x = y \ x < y)" - by (auto simp add: word_less_def) + by (auto simp add: order_class.le_less) lemma mod_plus_cong: assumes 1: "(b::int) = b'" @@ -4523,17 +4611,15 @@ "1 + n \ (0::'a::len word) \ unat (1 + n) = Suc (unat n)" by unat_arith - lemma word_no_1 [simp]: "(Numeral1::'a::len0 word) = 1" by (fact word_1_no [symmetric]) -lemma word_no_0 [simp]: "(Numeral0::'a::len0 word) = 0" - by (fact word_0_no [symmetric]) - declare bin_to_bl_def [simp] use "~~/src/HOL/Word/Tools/smt_word.ML" setup {* SMT_Word.setup *} +hide_const (open) Word + end diff -r 29e92b644d6c -r f067afe98049 src/HOL/ex/Arith_Examples.thy --- a/src/HOL/ex/Arith_Examples.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/ex/Arith_Examples.thy Mon Mar 26 10:56:56 2012 +0200 @@ -218,10 +218,10 @@ lemma "(0::int) < 1" by linarith -lemma "(47::nat) + 11 < 08 * 15" +lemma "(47::nat) + 11 < 8 * 15" by linarith -lemma "(47::int) + 11 < 08 * 15" +lemma "(47::int) + 11 < 8 * 15" by linarith text {* Splitting of inequalities of different type. *} diff -r 29e92b644d6c -r f067afe98049 src/HOL/ex/Code_Nat_examples.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/Code_Nat_examples.thy Mon Mar 26 10:56:56 2012 +0200 @@ -0,0 +1,57 @@ +(* Title: HOL/ex/Code_Nat_examples.thy + Author: Florian Haftmann, TU Muenchen +*) + +header {* Simple examples for Code\_Numeral\_Nat theory. *} + +theory Code_Nat_examples +imports Complex_Main "~~/src/HOL/Library/Efficient_Nat" +begin + +fun to_n :: "nat \ nat list" +where + "to_n 0 = []" +| "to_n (Suc 0) = []" +| "to_n (Suc (Suc 0)) = []" +| "to_n (Suc n) = n # to_n n" + +definition naive_prime :: "nat \ bool" +where + "naive_prime n \ n \ 2 \ filter (\m. n mod m = 0) (to_n n) = []" + +primrec fac :: "nat \ nat" +where + "fac 0 = 1" +| "fac (Suc n) = Suc n * fac n" + +primrec harmonic :: "nat \ rat" +where + "harmonic 0 = 0" +| "harmonic (Suc n) = 1 / of_nat (Suc n) + harmonic n" + +lemma "harmonic 200 \ 5" + by eval + +lemma "(let (q, r) = quotient_of (harmonic 8) in q div r) \ 2" + by normalization + +lemma "naive_prime 89" + by eval + +lemma "naive_prime 89" + by normalization + +lemma "\ naive_prime 87" + by eval + +lemma "\ naive_prime 87" + by normalization + +lemma "fac 10 > 3000000" + by eval + +lemma "fac 10 > 3000000" + by normalization + +end + diff -r 29e92b644d6c -r f067afe98049 src/HOL/ex/Dedekind_Real.thy --- a/src/HOL/ex/Dedekind_Real.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/ex/Dedekind_Real.thy Mon Mar 26 10:56:56 2012 +0200 @@ -1658,19 +1658,6 @@ by (blast intro!: real_less_all_preal linorder_not_less [THEN iffD1]) -subsection{*Numerals and Arithmetic*} - -instantiation real :: number_ring -begin - -definition - real_number_of_def: "(number_of w :: real) = of_int w" - -instance - by intro_classes (simp add: real_number_of_def) - -end - subsection {* Completeness of Positive Reals *} text {* diff -r 29e92b644d6c -r f067afe98049 src/HOL/ex/Efficient_Nat_examples.thy --- a/src/HOL/ex/Efficient_Nat_examples.thy Fri Mar 23 20:32:43 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,56 +0,0 @@ -(* Title: HOL/ex/Efficient_Nat_examples.thy - Author: Florian Haftmann, TU Muenchen -*) - -header {* Simple examples for Efficient\_Nat theory. *} - -theory Efficient_Nat_examples -imports Complex_Main "~~/src/HOL/Library/Efficient_Nat" -begin - -fun to_n :: "nat \ nat list" where - "to_n 0 = []" - | "to_n (Suc 0) = []" - | "to_n (Suc (Suc 0)) = []" - | "to_n (Suc n) = n # to_n n" - -definition naive_prime :: "nat \ bool" where - "naive_prime n \ n \ 2 \ filter (\m. n mod m = 0) (to_n n) = []" - -primrec fac :: "nat \ nat" where - "fac 0 = 1" - | "fac (Suc n) = Suc n * fac n" - -primrec rat_of_nat :: "nat \ rat" where - "rat_of_nat 0 = 0" - | "rat_of_nat (Suc n) = rat_of_nat n + 1" - -primrec harmonic :: "nat \ rat" where - "harmonic 0 = 0" - | "harmonic (Suc n) = 1 / rat_of_nat (Suc n) + harmonic n" - -lemma "harmonic 200 \ 5" - by eval - -lemma "harmonic 20 \ 3" - by normalization - -lemma "naive_prime 89" - by eval - -lemma "naive_prime 89" - by normalization - -lemma "\ naive_prime 87" - by eval - -lemma "\ naive_prime 87" - by normalization - -lemma "fac 10 > 3000000" - by eval - -lemma "fac 10 > 3000000" - by normalization - -end diff -r 29e92b644d6c -r f067afe98049 src/HOL/ex/Executable_Relation.thy --- a/src/HOL/ex/Executable_Relation.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/ex/Executable_Relation.thy Mon Mar 26 10:56:56 2012 +0200 @@ -12,6 +12,7 @@ "(x, y) : (rel_raw X R) = ((x = y \ x : X) \ (x, y) : R)" unfolding rel_raw_def by auto + lemma Id_raw: "Id = rel_raw UNIV {}" unfolding rel_raw_def by auto @@ -74,36 +75,34 @@ lemmas rel_raw_of_set_eqI[intro!] = arg_cong[where f="rel_of_set"] -definition rel :: "'a set => ('a * 'a) set => 'a rel" -where - "rel X R = rel_of_set (rel_raw X R)" +quotient_definition rel where "rel :: 'a set => ('a * 'a) set => 'a rel" is rel_raw done subsubsection {* Constant definitions on relations *} hide_const (open) converse rel_comp rtrancl Image quotient_definition member :: "'a * 'a => 'a rel => bool" where - "member" is "Set.member :: 'a * 'a => ('a * 'a) set => bool" + "member" is "Set.member :: 'a * 'a => ('a * 'a) set => bool" done quotient_definition converse :: "'a rel => 'a rel" where - "converse" is "Relation.converse :: ('a * 'a) set => ('a * 'a) set" + "converse" is "Relation.converse :: ('a * 'a) set => ('a * 'a) set" done quotient_definition union :: "'a rel => 'a rel => 'a rel" where - "union" is "Set.union :: ('a * 'a) set => ('a * 'a) set => ('a * 'a) set" + "union" is "Set.union :: ('a * 'a) set => ('a * 'a) set => ('a * 'a) set" done quotient_definition rel_comp :: "'a rel => 'a rel => 'a rel" where - "rel_comp" is "Relation.rel_comp :: ('a * 'a) set => ('a * 'a) set => ('a * 'a) set" + "rel_comp" is "Relation.rel_comp :: ('a * 'a) set => ('a * 'a) set => ('a * 'a) set" done quotient_definition rtrancl :: "'a rel => 'a rel" where - "rtrancl" is "Transitive_Closure.rtrancl :: ('a * 'a) set => ('a * 'a) set" + "rtrancl" is "Transitive_Closure.rtrancl :: ('a * 'a) set => ('a * 'a) set" done quotient_definition Image :: "'a rel => 'a set => 'a set" where - "Image" is "Relation.Image :: ('a * 'a) set => 'a set => 'a set" + "Image" is "Relation.Image :: ('a * 'a) set => 'a set => 'a set" done subsubsection {* Code generation *} @@ -111,33 +110,27 @@ lemma [code]: "member (x, y) (rel X R) = ((x = y \ x : X) \ (x, y) : R)" -unfolding rel_def member_def -by (simp add: member_raw) +by (lifting member_raw) lemma [code]: "converse (rel X R) = rel X (R^-1)" -unfolding rel_def converse_def -by (simp add: converse_raw) +by (lifting converse_raw) lemma [code]: "union (rel X R) (rel Y S) = rel (X Un Y) (R Un S)" -unfolding rel_def union_def -by (simp add: union_raw) +by (lifting union_raw) lemma [code]: "rel_comp (rel X R) (rel Y S) = rel (X Int Y) (Set.project (%(x, y). y : Y) R Un (Set.project (%(x, y). x : X) S Un R O S))" -unfolding rel_def rel_comp_def -by (simp add: rel_comp_raw) +by (lifting rel_comp_raw) lemma [code]: "rtrancl (rel X R) = rel UNIV (R^+)" -unfolding rel_def rtrancl_def -by (simp add: rtrancl_raw) +by (lifting rtrancl_raw) lemma [code]: "Image (rel X R) S = (X Int S) Un (R `` S)" -unfolding rel_def Image_def -by (simp add: Image_raw) +by (lifting Image_raw) quickcheck_generator rel constructors: rel diff -r 29e92b644d6c -r f067afe98049 src/HOL/ex/Groebner_Examples.thy --- a/src/HOL/ex/Groebner_Examples.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/ex/Groebner_Examples.thy Mon Mar 26 10:56:56 2012 +0200 @@ -31,7 +31,7 @@ (Conv.arg_conv (Conv.arg1_conv (Semiring_Normalizer.semiring_normalize_conv @{context})))) *}) by (rule refl) -lemma "((-3) ^ (Suc (Suc (Suc 0)))) == (X::'a::{number_ring})" +lemma "((-3) ^ (Suc (Suc (Suc 0)))) == (X::'a::{comm_ring_1})" apply (simp only: power_Suc power_0) apply (simp only: semiring_norm) oops @@ -58,7 +58,7 @@ by algebra lemma - fixes x::"'a::{idom,number_ring}" + fixes x::"'a::{idom}" shows "x^2*y = x^2 & x*y^2 = y^2 \ x=1 & y=1 | x=0 & y=0" by algebra @@ -69,7 +69,7 @@ "sq x == x*x" lemma - fixes x1 :: "'a::{idom,number_ring}" + fixes x1 :: "'a::{idom}" shows "(sq x1 + sq x2 + sq x3 + sq x4) * (sq y1 + sq y2 + sq y3 + sq y4) = sq (x1*y1 - x2*y2 - x3*y3 - x4*y4) + @@ -79,7 +79,7 @@ by (algebra add: sq_def) lemma - fixes p1 :: "'a::{idom,number_ring}" + fixes p1 :: "'a::{idom}" shows "(sq p1 + sq q1 + sq r1 + sq s1 + sq t1 + sq u1 + sq v1 + sq w1) * (sq p2 + sq q2 + sq r2 + sq s2 + sq t2 + sq u2 + sq v2 + sq w2) diff -r 29e92b644d6c -r f067afe98049 src/HOL/ex/Numeral_Representation.thy --- a/src/HOL/ex/Numeral_Representation.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/ex/Numeral_Representation.thy Mon Mar 26 10:56:56 2012 +0200 @@ -2,7 +2,7 @@ Author: Florian Haftmann *) -header {* An experimental alternative numeral representation. *} +header {* First experiments for a numeral representation (now obsolete). *} theory Numeral_Representation imports Main @@ -498,7 +498,7 @@ by (simp add: less_imp_le minus_of_num_less_one_iff) lemma minus_one_le_of_num_iff: "- 1 \ of_num n" - by (simp add: less_imp_le minus_one_less_of_num_iff) + by (simp only: less_imp_le minus_one_less_of_num_iff) lemma minus_one_le_one_iff: "- 1 \ 1" by (simp add: less_imp_le minus_one_less_one_iff) @@ -510,7 +510,7 @@ by (simp add: not_le minus_of_num_less_one_iff) lemma of_num_le_minus_one_iff: "\ of_num n \ - 1" - by (simp add: not_le minus_one_less_of_num_iff) + by (simp only: not_le minus_one_less_of_num_iff) lemma one_le_minus_one_iff: "\ 1 \ - 1" by (simp add: not_le minus_one_less_one_iff) @@ -522,10 +522,10 @@ by (simp add: not_less minus_of_num_le_one_iff) lemma of_num_less_minus_one_iff: "\ of_num n < - 1" - by (simp add: not_less minus_one_le_of_num_iff) + by (simp only: not_less minus_one_le_of_num_iff) lemma one_less_minus_one_iff: "\ 1 < - 1" - by (simp add: not_less minus_one_le_one_iff) + by (simp only: not_less minus_one_le_one_iff) lemmas le_signed_numeral_special [numeral] = minus_of_num_le_of_num_iff @@ -835,10 +835,7 @@ text {* Reversing standard setup *} -lemma [code_unfold del]: "(0::int) \ Numeral0" by simp lemma [code_unfold del]: "(1::int) \ Numeral1" by simp -declare zero_is_num_zero [code_unfold del] -declare one_is_num_one [code_unfold del] lemma [code, code del]: "(1 :: int) = 1" @@ -970,147 +967,5 @@ hide_const (open) sub dup -text {* Pretty literals *} - -ML {* -local open Code_Thingol in - -fun add_code print target = - let - fun dest_num one' dig0' dig1' thm = - let - fun dest_dig (IConst (c, _)) = if c = dig0' then 0 - else if c = dig1' then 1 - else Code_Printer.eqn_error thm "Illegal numeral expression: illegal dig" - | dest_dig _ = Code_Printer.eqn_error thm "Illegal numeral expression: illegal digit"; - fun dest_num (IConst (c, _)) = if c = one' then 1 - else Code_Printer.eqn_error thm "Illegal numeral expression: illegal leading digit" - | dest_num (t1 `$ t2) = 2 * dest_num t2 + dest_dig t1 - | dest_num _ = Code_Printer.eqn_error thm "Illegal numeral expression: illegal term"; - in dest_num end; - fun pretty sgn literals [one', dig0', dig1'] _ thm _ _ [(t, _)] = - (Code_Printer.str o print literals o sgn o dest_num one' dig0' dig1' thm) t - fun add_syntax (c, sgn) = Code_Target.add_const_syntax target c - (SOME (Code_Printer.complex_const_syntax - (1, ([@{const_name One}, @{const_name Dig0}, @{const_name Dig1}], - pretty sgn)))); - in - add_syntax (@{const_name Pls}, I) - #> add_syntax (@{const_name Mns}, (fn k => ~ k)) - end; - end -*} - -hide_const (open) One Dig0 Dig1 - - -subsection {* Toy examples *} - -definition "foo \ #4 * #2 + #7 = (#8 :: nat)" -definition "bar \ #4 * #2 + #7 \ (#8 :: int) - #3" - -code_thms foo bar -export_code foo bar checking SML OCaml? Haskell? Scala? - -text {* This is an ad-hoc @{text Code_Integer} setup. *} - -setup {* - fold (add_code Code_Printer.literal_numeral) - [Code_ML.target_SML, Code_ML.target_OCaml, Code_Haskell.target, Code_Scala.target] -*} - -code_type int - (SML "IntInf.int") - (OCaml "Big'_int.big'_int") - (Haskell "Integer") - (Scala "BigInt") - (Eval "int") - -code_const "0::int" - (SML "0/ :/ IntInf.int") - (OCaml "Big'_int.zero") - (Haskell "0") - (Scala "BigInt(0)") - (Eval "0/ :/ int") - -code_const Int.pred - (SML "IntInf.- ((_), 1)") - (OCaml "Big'_int.pred'_big'_int") - (Haskell "!(_/ -/ 1)") - (Scala "!(_ -/ 1)") - (Eval "!(_/ -/ 1)") -code_const Int.succ - (SML "IntInf.+ ((_), 1)") - (OCaml "Big'_int.succ'_big'_int") - (Haskell "!(_/ +/ 1)") - (Scala "!(_ +/ 1)") - (Eval "!(_/ +/ 1)") - -code_const "op + \ int \ int \ int" - (SML "IntInf.+ ((_), (_))") - (OCaml "Big'_int.add'_big'_int") - (Haskell infixl 6 "+") - (Scala infixl 7 "+") - (Eval infixl 8 "+") - -code_const "uminus \ int \ int" - (SML "IntInf.~") - (OCaml "Big'_int.minus'_big'_int") - (Haskell "negate") - (Scala "!(- _)") - (Eval "~/ _") - -code_const "op - \ int \ int \ int" - (SML "IntInf.- ((_), (_))") - (OCaml "Big'_int.sub'_big'_int") - (Haskell infixl 6 "-") - (Scala infixl 7 "-") - (Eval infixl 8 "-") - -code_const "op * \ int \ int \ int" - (SML "IntInf.* ((_), (_))") - (OCaml "Big'_int.mult'_big'_int") - (Haskell infixl 7 "*") - (Scala infixl 8 "*") - (Eval infixl 9 "*") - -code_const pdivmod - (SML "IntInf.divMod/ (IntInf.abs _,/ IntInf.abs _)") - (OCaml "Big'_int.quomod'_big'_int/ (Big'_int.abs'_big'_int _)/ (Big'_int.abs'_big'_int _)") - (Haskell "divMod/ (abs _)/ (abs _)") - (Scala "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))") - (Eval "Integer.div'_mod/ (abs _)/ (abs _)") - -code_const "HOL.equal \ int \ int \ bool" - (SML "!((_ : IntInf.int) = _)") - (OCaml "Big'_int.eq'_big'_int") - (Haskell infix 4 "==") - (Scala infixl 5 "==") - (Eval infixl 6 "=") - -code_const "op \ \ int \ int \ bool" - (SML "IntInf.<= ((_), (_))") - (OCaml "Big'_int.le'_big'_int") - (Haskell infix 4 "<=") - (Scala infixl 4 "<=") - (Eval infixl 6 "<=") - -code_const "op < \ int \ int \ bool" - (SML "IntInf.< ((_), (_))") - (OCaml "Big'_int.lt'_big'_int") - (Haskell infix 4 "<") - (Scala infixl 4 "<") - (Eval infixl 6 "<") - -code_const Code_Numeral.int_of - (SML "IntInf.fromInt") - (OCaml "_") - (Haskell "toInteger") - (Scala "!_.as'_BigInt") - (Eval "_") - -export_code foo bar checking SML OCaml? Haskell? Scala? - -end diff -r 29e92b644d6c -r f067afe98049 src/HOL/ex/ROOT.ML --- a/src/HOL/ex/ROOT.ML Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/ex/ROOT.ML Mon Mar 26 10:56:56 2012 +0200 @@ -5,7 +5,7 @@ no_document use_thys [ "~~/src/HOL/Library/State_Monad", - "Efficient_Nat_examples", + "Code_Nat_examples", "~~/src/HOL/Library/FuncSet", "Eval_Examples", "Normalization_by_Evaluation", diff -r 29e92b644d6c -r f067afe98049 src/HOL/ex/ReflectionEx.thy --- a/src/HOL/ex/ReflectionEx.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/ex/ReflectionEx.thy Mon Mar 26 10:56:56 2012 +0200 @@ -143,7 +143,7 @@ oops text{* Hmmm let's specialize @{text Inum_C} with numerals.*} -lemma Inum_number: "Inum (C (number_of t)) vs = number_of t" by simp +lemma Inum_number: "Inum (C (numeral t)) vs = numeral t" by simp lemmas Inum_eqs = Inum_Var Inum_Add Inum_Mul Inum_CN Inum_number text{* Second attempt *} @@ -155,7 +155,7 @@ lemma "1 * (2* x + (y::nat) + 0 + 1) \ 0" apply (reify Inum_eqs ("1 * (2*x + (y::nat) + 0 + 1)")) oops - text{* Oh!! 0 is not a variable \dots\ Oh! 0 is not a @{text "number_of"} \dots\ thing. The same for 1. So let's add those equations too *} + text{* Oh!! 0 is not a variable \dots\ Oh! 0 is not a @{text "numeral"} \dots\ thing. The same for 1. So let's add those equations too *} lemma Inum_01: "Inum (C 0) vs = 0" "Inum (C 1) vs = 1" "Inum (C(Suc n)) vs = Suc n" by simp+ @@ -312,9 +312,9 @@ by simp lemma Irint_C1: "Irint (IC 1) vs = 1" by simp -lemma Irint_Cnumberof: "Irint (IC (number_of x)) vs = number_of x" +lemma Irint_Cnumeral: "Irint (IC (numeral x)) vs = numeral x" by simp -lemmas Irint_simps = Irint_Var Irint_Neg Irint_Add Irint_Sub Irint_Mult Irint_C0 Irint_C1 Irint_Cnumberof +lemmas Irint_simps = Irint_Var Irint_Neg Irint_Add Irint_Sub Irint_Mult Irint_C0 Irint_C1 Irint_Cnumeral lemma "(3::int) * x + y*y - 9 + (- z) = 0" apply (reify Irint_simps ("(3::int) * x + y*y - 9 + (- z)")) oops @@ -348,10 +348,10 @@ by simp lemma Irnat_C1: "Irnat (NC 1) is ls vs = 1" by simp -lemma Irnat_Cnumberof: "Irnat (NC (number_of x)) is ls vs = number_of x" +lemma Irnat_Cnumeral: "Irnat (NC (numeral x)) is ls vs = numeral x" by simp lemmas Irnat_simps = Irnat_Suc Irnat_Var Irnat_Neg Irnat_Add Irnat_Sub Irnat_Mult Irnat_lgth - Irnat_C0 Irnat_C1 Irnat_Cnumberof + Irnat_C0 Irnat_C1 Irnat_Cnumeral lemma "(Suc n) * length (([(3::int) * x + y*y - 9 + (- z)] @ []) @ xs) = length xs" apply (reify Irnat_simps Irlist.simps Irint_simps ("(Suc n) *length (([(3::int) * x + y*y - 9 + (- z)] @ []) @ xs)")) oops diff -r 29e92b644d6c -r f067afe98049 src/HOL/ex/Simproc_Tests.thy --- a/src/HOL/ex/Simproc_Tests.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/HOL/ex/Simproc_Tests.thy Mon Mar 26 10:56:56 2012 +0200 @@ -5,7 +5,7 @@ header {* Testing of arithmetic simprocs *} theory Simproc_Tests -imports Main +imports (*Main*) "../Numeral_Simprocs" begin text {* @@ -43,7 +43,7 @@ possible. *) notepad begin - fix a b c d oo uu i j k l u v w x y z :: "'a::number_ring" + fix a b c d oo uu i j k l u v w x y z :: "'a::comm_ring_1" { assume "a + - b = u" have "(a + c) - (b + c) = u" by (tactic {* test [@{simproc int_combine_numerals}] *}) fact @@ -107,7 +107,7 @@ subsection {* @{text inteq_cancel_numerals} *} notepad begin - fix i j k u vv w y z w' y' z' :: "'a::number_ring" + fix i j k u vv w y z w' y' z' :: "'a::comm_ring_1" { assume "u = 0" have "2*u = u" by (tactic {* test [@{simproc inteq_cancel_numerals}] *}) fact @@ -130,7 +130,7 @@ subsection {* @{text intless_cancel_numerals} *} notepad begin - fix b c i j k u y :: "'a::{linordered_idom,number_ring}" + fix b c i j k u y :: "'a::linordered_idom" { assume "y < 2 * b" have "y - b < b" by (tactic {* test [@{simproc intless_cancel_numerals}] *}) fact @@ -151,7 +151,7 @@ subsection {* @{text ring_eq_cancel_numeral_factor} *} notepad begin - fix x y :: "'a::{idom,ring_char_0,number_ring}" + fix x y :: "'a::{idom,ring_char_0}" { assume "3*x = 4*y" have "9*x = 12 * y" by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact @@ -176,7 +176,7 @@ subsection {* @{text int_div_cancel_numeral_factors} *} notepad begin - fix x y z :: "'a::{semiring_div,ring_char_0,number_ring}" + fix x y z :: "'a::{semiring_div,comm_ring_1,ring_char_0}" { assume "(3*x) div (4*y) = z" have "(9*x) div (12*y) = z" by (tactic {* test [@{simproc int_div_cancel_numeral_factors}] *}) fact @@ -199,7 +199,7 @@ subsection {* @{text ring_less_cancel_numeral_factor} *} notepad begin - fix x y :: "'a::{linordered_idom,number_ring}" + fix x y :: "'a::linordered_idom" { assume "3*x < 4*y" have "9*x < 12 * y" by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact @@ -224,7 +224,7 @@ subsection {* @{text ring_le_cancel_numeral_factor} *} notepad begin - fix x y :: "'a::{linordered_idom,number_ring}" + fix x y :: "'a::linordered_idom" { assume "3*x \ 4*y" have "9*x \ 12 * y" by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact @@ -255,7 +255,7 @@ subsection {* @{text divide_cancel_numeral_factor} *} notepad begin - fix x y z :: "'a::{field_inverse_zero,ring_char_0,number_ring}" + fix x y z :: "'a::{field_inverse_zero,ring_char_0}" { assume "(3*x) / (4*y) = z" have "(9*x) / (12 * y) = z" by (tactic {* test [@{simproc divide_cancel_numeral_factor}] *}) fact @@ -322,6 +322,9 @@ } end +lemma shows "a*(b*c)/(y*z) = d*(b::'a::linordered_field_inverse_zero)*(x*a)/z" +oops -- "FIXME: need simproc to cover this case" + subsection {* @{text divide_cancel_factor} *} notepad begin @@ -393,7 +396,7 @@ subsection {* @{text field_combine_numerals} *} notepad begin - fix x y z uu :: "'a::{field_inverse_zero,ring_char_0,number_ring}" + fix x y z uu :: "'a::{field_inverse_zero,ring_char_0}" { assume "5 / 6 * x = uu" have "x / 2 + x / 3 = uu" by (tactic {* test [@{simproc field_combine_numerals}] *}) fact @@ -415,7 +418,7 @@ end lemma - fixes x :: "'a::{linordered_field_inverse_zero,number_ring}" + fixes x :: "'a::{linordered_field_inverse_zero}" shows "2/3 * x + x / 3 = uu" apply (tactic {* test [@{simproc field_combine_numerals}] *})? oops -- "FIXME: test fails" @@ -448,17 +451,12 @@ } end -(*negative numerals: FAIL*) -lemma "Suc (i + j + -3 + k) = u" -apply (tactic {* test [@{simproc nat_combine_numerals}] *})? -oops - subsection {* @{text nateq_cancel_numerals} *} notepad begin fix i j k l oo u uu vv w y z w' y' z' :: "nat" { - assume "Suc 0 * u = 0" have "2*u = u" + assume "Suc 0 * u = 0" have "2*u = (u::nat)" by (tactic {* test [@{simproc nateq_cancel_numerals}] *}) fact next assume "Suc 0 * u = Suc 0" have "2*u = Suc (u)" @@ -504,7 +502,7 @@ notepad begin fix length :: "'a \ nat" and l1 l2 xs :: "'a" and f :: "nat \ 'a" - fix c i j k l oo u uu vv w y z w' y' z' :: "nat" + fix c i j k l m oo u uu vv w y z w' y' z' :: "nat" { assume "0 < j" have "(2*length xs < 2*length xs + j)" by (tactic {* test [@{simproc natless_cancel_numerals}] *}) fact @@ -518,14 +516,6 @@ next assume "0 < Suc 0 * (m * n) + u" have "(2*n*m) < (3*(m*n)) + u" by (tactic {* test [@{simproc natless_cancel_numerals}] *}) fact - next - (* FIXME: negative numerals fail - have "(i + j + -23 + (k::nat)) < u + 15 + y" - apply (tactic {* test [@{simproc natless_cancel_numerals}] *})? - sorry - have "(i + j + 3 + (k::nat)) < u + -15 + y" - apply (tactic {* test [@{simproc natless_cancel_numerals}] *})? - sorry*) } end @@ -611,17 +601,6 @@ next assume "u + y - 0 = v" have "Suc (Suc (Suc (Suc (Suc (u + y))))) - 5 = v" by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact - next - (* FIXME: negative numerals fail - have "(i + j + -12 + k) - 15 = y" - apply (tactic {* test [@{simproc natdiff_cancel_numerals}] *})? - sorry - have "(i + j + 12 + k) - -15 = y" - apply (tactic {* test [@{simproc natdiff_cancel_numerals}] *})? - sorry - have "(i + j + -12 + k) - -15 = y" - apply (tactic {* test [@{simproc natdiff_cancel_numerals}] *})? - sorry*) } end diff -r 29e92b644d6c -r f067afe98049 src/ZF/AC/AC_Equiv.thy --- a/src/ZF/AC/AC_Equiv.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/ZF/AC/AC_Equiv.thy Mon Mar 26 10:56:56 2012 +0200 @@ -162,11 +162,6 @@ "[| f \ inj(A, B); !!a. a \ A ==> f`a \ C |] ==> f \ inj(A,C)" by (unfold inj_def, blast intro: Pi_type) -lemma nat_not_Finite: "~ Finite(nat)" -by (unfold Finite_def, blast dest: eqpoll_imp_lepoll ltI lt_not_lepoll) - -lemmas le_imp_lepoll = le_imp_subset [THEN subset_imp_lepoll] - (* ********************************************************************** *) (* Another elimination rule for \! *) (* ********************************************************************** *) @@ -175,30 +170,18 @@ by blast (* ********************************************************************** *) -(* image of a surjection *) -(* ********************************************************************** *) - -lemma surj_image_eq: "f \ surj(A, B) ==> f``A = B" -apply (unfold surj_def) -apply (erule CollectE) -apply (rule image_fun [THEN ssubst], assumption, rule subset_refl) -apply (blast dest: apply_type) -done - - -(* ********************************************************************** *) (* Lemmas used in the proofs like WO? ==> AC? *) (* ********************************************************************** *) lemma first_in_B: - "[| well_ord(\(A),r); 0\A; B \ A |] ==> (THE b. first(b,B,r)) \ B" + "[| well_ord(\(A),r); 0 \ A; B \ A |] ==> (THE b. first(b,B,r)) \ B" by (blast dest!: well_ord_imp_ex1_first [THEN theI, THEN first_def [THEN def_imp_iff, THEN iffD1]]) -lemma ex_choice_fun: "[| well_ord(\(A), R); 0\A |] ==> \f. f:(\ X \ A. X)" +lemma ex_choice_fun: "[| well_ord(\(A), R); 0 \ A |] ==> \f. f \ (\ X \ A. X)" by (fast elim!: first_in_B intro!: lam_type) -lemma ex_choice_fun_Pow: "well_ord(A, R) ==> \f. f:(\ X \ Pow(A)-{0}. X)" +lemma ex_choice_fun_Pow: "well_ord(A, R) ==> \f. f \ (\ X \ Pow(A)-{0}. X)" by (fast elim!: well_ord_subset [THEN ex_choice_fun]) diff -r 29e92b644d6c -r f067afe98049 src/ZF/AC/Cardinal_aux.thy --- a/src/ZF/AC/Cardinal_aux.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/ZF/AC/Cardinal_aux.thy Mon Mar 26 10:56:56 2012 +0200 @@ -30,46 +30,32 @@ "[| A \ i; Ord(i) |] ==> \j. j j" by (unfold lesspoll_def, blast dest!: lepoll_imp_ex_le_eqpoll elim!: leE) -lemma Inf_Ord_imp_InfCard_cardinal: "[| ~Finite(i); Ord(i) |] ==> InfCard(|i|)" -apply (unfold InfCard_def) -apply (rule conjI) -apply (rule Card_cardinal) -apply (rule Card_nat - [THEN Card_def [THEN def_imp_iff, THEN iffD1, THEN ssubst]]) - -- "rewriting would loop!" -apply (rule well_ord_Memrel [THEN well_ord_lepoll_imp_Card_le], assumption) -apply (rule nat_le_infinite_Ord [THEN le_imp_lepoll], assumption+) -done - -text{*An alternative and more general proof goes like this: A and B are both -well-ordered (because they are injected into an ordinal), either @{term"A \ B"} -or @{term"B \ A"}. Also both are equipollent to their cardinalities, so -(if A and B are infinite) then @{term"A \ B \ |A\B| \ max(|A|,|B|) \ i"}. -In fact, the correctly strengthened version of this theorem appears below.*} -lemma Un_lepoll_Inf_Ord_weak: - "[|A \ i; B \ i; \ Finite(i); Ord(i)|] ==> A \ B \ i" -apply (rule Un_lepoll_sum [THEN lepoll_trans]) -apply (rule lepoll_imp_sum_lepoll_prod [THEN lepoll_trans]) -apply (erule eqpoll_trans [THEN eqpoll_imp_lepoll]) -apply (erule eqpoll_sym) -apply (rule subset_imp_lepoll [THEN lepoll_trans, THEN lepoll_trans]) -apply (rule nat_2I [THEN OrdmemD], rule Ord_nat) -apply (rule nat_le_infinite_Ord [THEN le_imp_lepoll], assumption+) -apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll]) -apply (erule prod_eqpoll_cong [THEN eqpoll_imp_lepoll, THEN lepoll_trans], - assumption) -apply (rule eqpoll_imp_lepoll) -apply (rule well_ord_Memrel [THEN well_ord_InfCard_square_eq], assumption) -apply (rule Inf_Ord_imp_InfCard_cardinal, assumption+) -done - lemma Un_eqpoll_Inf_Ord: - "[| A \ i; B \ i; ~Finite(i); Ord(i) |] ==> A \ B \ i" -apply (rule eqpollI) -apply (blast intro: Un_lepoll_Inf_Ord_weak) -apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll, THEN lepoll_trans]) -apply (rule Un_upper1 [THEN subset_imp_lepoll]) -done + assumes A: "A \ i" and B: "B \ i" and NFI: "\ Finite(i)" and i: "Ord(i)" + shows "A \ B \ i" +proof (rule eqpollI) + have AB: "A \ B" using A B by (blast intro: eqpoll_sym eqpoll_trans) + have "2 \ nat" + by (rule subset_imp_lepoll) (rule OrdmemD [OF nat_2I Ord_nat]) + also have "... \ i" + by (simp add: nat_le_infinite_Ord le_imp_lepoll NFI i)+ + also have "... \ A" by (blast intro: eqpoll_sym A) + finally have "2 \ A" . + have ICI: "InfCard(|i|)" + by (simp add: Inf_Card_is_InfCard Finite_cardinal_iff NFI i) + have "A \ B \ A + B" by (rule Un_lepoll_sum) + also have "... \ A \ B" + by (rule lepoll_imp_sum_lepoll_prod [OF AB [THEN eqpoll_imp_lepoll] `2 \ A`]) + also have "... \ i \ i" + by (blast intro: prod_eqpoll_cong eqpoll_imp_lepoll A B) + also have "... \ i" + by (blast intro: well_ord_InfCard_square_eq well_ord_Memrel ICI i) + finally show "A \ B \ i" . +next + have "i \ A" by (blast intro: A eqpoll_sym) + also have "... \ A \ B" by (blast intro: subset_imp_lepoll) + finally show "i \ A \ B" . +qed schematic_lemma paired_bij: "?f \ bij({{y,z}. y \ x}, x)" apply (rule RepFun_bijective) @@ -82,8 +68,7 @@ lemma ex_eqpoll_disjoint: "\B. B \ A & B \ C = 0" by (fast intro!: paired_eqpoll equals0I elim: mem_asym) -(*Finally we reach this result. Surely there's a simpler proof, as sketched - above?*) +(*Finally we reach this result. Surely there's a simpler proof?*) lemma Un_lepoll_Inf_Ord: "[| A \ i; B \ i; ~Finite(i); Ord(i) |] ==> A \ B \ i" apply (rule_tac A1 = i and C1 = i in ex_eqpoll_disjoint [THEN exE]) diff -r 29e92b644d6c -r f067afe98049 src/ZF/Cardinal.thy --- a/src/ZF/Cardinal.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/ZF/Cardinal.thy Mon Mar 26 10:56:56 2012 +0200 @@ -445,7 +445,7 @@ (*Infinite unions of cardinals? See Devlin, Lemma 6.7, page 98*) -lemma Card_cardinal: "Card(|A|)" +lemma Card_cardinal [iff]: "Card(|A|)" proof (unfold cardinal_def) show "Card(\ i. i \ A)" proof (cases "\i. Ord (i) & i \ A") @@ -1105,6 +1105,9 @@ lemma Finite_Pow_iff [iff]: "Finite(Pow(A)) \ Finite(A)" by (blast intro: Finite_Pow Finite_Pow_imp_Finite) +lemma Finite_cardinal_iff: + assumes i: "Ord(i)" shows "Finite(|i|) \ Finite(i)" + by (auto simp add: Finite_def) (blast intro: eqpoll_trans eqpoll_sym Ord_cardinal_eqpoll [OF i])+ (*Krzysztof Grabczewski's proof that the converse of a finite, well-ordered diff -r 29e92b644d6c -r f067afe98049 src/ZF/CardinalArith.thy --- a/src/ZF/CardinalArith.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/ZF/CardinalArith.thy Mon Mar 26 10:56:56 2012 +0200 @@ -682,7 +682,7 @@ apply (simp add: InfCard_is_Card [THEN Card_cardinal_eq]) done -lemma Inf_Card_is_InfCard: "[| ~Finite(i); Card(i) |] ==> InfCard(i)" +lemma Inf_Card_is_InfCard: "[| Card(i); ~ Finite(i) |] ==> InfCard(i)" by (simp add: InfCard_def Card_is_Ord [THEN nat_le_infinite_Ord]) subsubsection{*Toward's Kunen's Corollary 10.13 (1)*} diff -r 29e92b644d6c -r f067afe98049 src/ZF/Perm.thy --- a/src/ZF/Perm.thy Fri Mar 23 20:32:43 2012 +0100 +++ b/src/ZF/Perm.thy Mon Mar 26 10:56:56 2012 +0200 @@ -505,6 +505,9 @@ apply (blast intro: apply_equality apply_Pair Pi_type) done +lemma surj_image_eq: "f \ surj(A, B) ==> f``A = B" + by (auto simp add: surj_def image_fun) (blast dest: apply_type) + lemma restrict_image [simp]: "restrict(f,A) `` B = f `` (A \ B)" by (auto simp add: restrict_def)