src/HOL/ex/Hex_Bin_Examples.thy
author krauss
Tue, 28 Sep 2010 09:54:07 +0200
changeset 39754 150f831ce4a3
parent 20866 bc366b4b6ea4
child 41460 ea56b98aee83
permissions -rw-r--r--
no longer declare .psimps rules as [simp]. This regularly caused confusion (e.g., they show up in simp traces when the regular simp rules are disabled). In the few places where the rules are used, explicitly mentioning them actually clarifies the proof text.

(*  Title:      HOL/ex/Hex_Bin_Examples.thy
    ID:         $Id$
    Author:     Gerwin Klein, NICTA
*)

header {* Examples for hexadecimal and binary numerals *}

theory Hex_Bin_Examples imports Main 
begin


text "Hex and bin numerals can be used like normal decimal numerals in input"
lemma "0xFF = 255" by (rule refl)
lemma "0xF = 0b1111" by (rule refl)

text {* 
  Just like decimal numeral they are polymorphic, for arithmetic 
  they need to be constrained
*}
lemma "0x0A + 0x10 = (0x1A :: nat)" by simp

text "The number of leading zeros is irrelevant"
lemma "0b00010000 = 0x10" by (rule refl) 

text "Unary minus works as for decimal numerals"
lemma "- 0x0A = - 10" by (rule refl)

text {*
  Hex and bin numerals are printed as decimal: @{term "0b10"}
*}
term "0b10"
term "0x0A"

text {* 
  The numerals 0 and 1 are syntactically different from the 
  constants 0 and 1. For the usual numeric types, their values 
  are the same, though.
*}
lemma "0x01 = 1" oops 
lemma "0x00 = 0" oops 

lemma "0x01 = (1::nat)" by simp
lemma "0b0000 = (0::int)" by simp


end