src/HOL/Word/Bit_Representation.thy
changeset 37667 41acc0fa6b6c
parent 37658 df789294c77a
child 41413 64cd30d6b0b8
equal deleted inserted replaced
37663:f2c98b8c0c5c 37667:41acc0fa6b6c
     1 (* 
     1 (* 
     2   Author: Jeremy Dawson, NICTA
     2   Author: Jeremy Dawson, NICTA
     3 
     3 
     4   contains basic definition to do with integers
     4   contains basic definition to do with integers
     5   expressed using Pls, Min, BIT and important resulting theorems, 
     5   expressed using Pls, Min, BIT
     6   in particular, bin_rec and related work
       
     7 *) 
     6 *) 
     8 
     7 
     9 header {* Basic Definitions for Binary Integers *}
     8 header {* Basic Definitions for Binary Integers *}
    10 
     9 
    11 theory Bit_Representation
    10 theory Bit_Representation
    12 imports Misc_Numeric Bit
    11 imports Misc_Numeric Bit
    13 begin
    12 begin
    14 
    13 
    15 subsection {* Further properties of numerals *}
    14 subsection {* Further properties of numerals *}
    16 
    15 
       
    16 definition bitval :: "bit \<Rightarrow> 'a\<Colon>zero_neq_one" where
       
    17   "bitval = bit_case 0 1"
       
    18 
       
    19 lemma bitval_simps [simp]:
       
    20   "bitval 0 = 0"
       
    21   "bitval 1 = 1"
       
    22   by (simp_all add: bitval_def)
       
    23 
    17 definition Bit :: "int \<Rightarrow> bit \<Rightarrow> int" (infixl "BIT" 90) where
    24 definition Bit :: "int \<Rightarrow> bit \<Rightarrow> int" (infixl "BIT" 90) where
    18   "k BIT b = bit_case 0 1 b + k + k"
    25   "k BIT b = bitval b + k + k"
    19 
    26 
    20 lemma BIT_B0_eq_Bit0 [simp]: "w BIT 0 = Int.Bit0 w"
    27 lemma BIT_B0_eq_Bit0 [simp]: "w BIT 0 = Int.Bit0 w"
    21   unfolding Bit_def Bit0_def by simp
    28   unfolding Bit_def Bit0_def by simp
    22 
    29 
    23 lemma BIT_B1_eq_Bit1 [simp]: "w BIT 1 = Int.Bit1 w"
    30 lemma BIT_B1_eq_Bit1 [simp]: "w BIT 1 = Int.Bit1 w"
    41   by (rule FalseE)
    48   by (rule FalseE)
    42 
    49 
    43 (** ways in which type Bin resembles a datatype **)
    50 (** ways in which type Bin resembles a datatype **)
    44 
    51 
    45 lemma BIT_eq: "u BIT b = v BIT c ==> u = v & b = c"
    52 lemma BIT_eq: "u BIT b = v BIT c ==> u = v & b = c"
    46   apply (unfold Bit_def)
    53   apply (cases b) apply (simp_all)
    47   apply (simp (no_asm_use) split: bit.split_asm)
    54   apply (cases c) apply (simp_all)
    48      apply simp_all
    55   apply (cases c) apply (simp_all)
    49    apply (drule_tac f=even in arg_cong, clarsimp)+
       
    50   done
    56   done
    51      
    57      
    52 lemmas BIT_eqE [elim!] = BIT_eq [THEN conjE, standard]
    58 lemmas BIT_eqE [elim!] = BIT_eq [THEN conjE, standard]
    53 
    59 
    54 lemma BIT_eq_iff [simp]: 
    60 lemma BIT_eq_iff [simp]: 
    57 
    63 
    58 lemmas BIT_eqI [intro!] = conjI [THEN BIT_eq_iff [THEN iffD2]]
    64 lemmas BIT_eqI [intro!] = conjI [THEN BIT_eq_iff [THEN iffD2]]
    59 
    65 
    60 lemma less_Bits: 
    66 lemma less_Bits: 
    61   "(v BIT b < w BIT c) = (v < w | v <= w & b = (0::bit) & c = (1::bit))"
    67   "(v BIT b < w BIT c) = (v < w | v <= w & b = (0::bit) & c = (1::bit))"
    62   unfolding Bit_def by (auto split: bit.split)
    68   unfolding Bit_def by (auto simp add: bitval_def split: bit.split)
    63 
    69 
    64 lemma le_Bits: 
    70 lemma le_Bits: 
    65   "(v BIT b <= w BIT c) = (v < w | v <= w & (b ~= (1::bit) | c ~= (0::bit)))" 
    71   "(v BIT b <= w BIT c) = (v < w | v <= w & (b ~= (1::bit) | c ~= (0::bit)))" 
    66   unfolding Bit_def by (auto split: bit.split)
    72   unfolding Bit_def by (auto simp add: bitval_def split: bit.split)
    67 
    73 
    68 lemma no_no [simp]: "number_of (number_of i) = i"
    74 lemma no_no [simp]: "number_of (number_of i) = i"
    69   unfolding number_of_eq by simp
    75   unfolding number_of_eq by simp
    70 
    76 
    71 lemma Bit_B0:
    77 lemma Bit_B0:
   105 
   111 
   106 lemma bin_ex_rl: "EX w b. w BIT b = bin"
   112 lemma bin_ex_rl: "EX w b. w BIT b = bin"
   107   apply (unfold Bit_def)
   113   apply (unfold Bit_def)
   108   apply (cases "even bin")
   114   apply (cases "even bin")
   109    apply (clarsimp simp: even_equiv_def)
   115    apply (clarsimp simp: even_equiv_def)
   110    apply (auto simp: odd_equiv_def split: bit.split)
   116    apply (auto simp: odd_equiv_def bitval_def split: bit.split)
   111   done
   117   done
   112 
   118 
   113 lemma bin_exhaust:
   119 lemma bin_exhaust:
   114   assumes Q: "\<And>x b. bin = x BIT b \<Longrightarrow> Q"
   120   assumes Q: "\<And>x b. bin = x BIT b \<Longrightarrow> Q"
   115   shows "Q"
   121   shows "Q"
   235   apply (rule trans)
   241   apply (rule trans)
   236    apply clarsimp
   242    apply clarsimp
   237    apply (rule refl)
   243    apply (rule refl)
   238   apply (drule trans)
   244   apply (drule trans)
   239    apply (rule Bit_def)
   245    apply (rule Bit_def)
   240   apply (simp add: z1pdiv2 split: bit.split)
   246   apply (simp add: bitval_def z1pdiv2 split: bit.split)
   241   done
   247   done
   242 
   248 
   243 lemma Bit_div2 [simp]: "(w BIT b) div 2 = w"
   249 lemma Bit_div2 [simp]: "(w BIT b) div 2 = w"
   244   unfolding bin_rest_div [symmetric] by auto
   250   unfolding bin_rest_div [symmetric] by auto
   245 
   251 
   312 lemmas bin_nth_simps = 
   318 lemmas bin_nth_simps = 
   313   bin_nth_0 bin_nth_Suc bin_nth_Pls bin_nth_Min bin_nth_minus
   319   bin_nth_0 bin_nth_Suc bin_nth_Pls bin_nth_Min bin_nth_minus
   314   bin_nth_minus_Bit0 bin_nth_minus_Bit1
   320   bin_nth_minus_Bit0 bin_nth_minus_Bit1
   315 
   321 
   316 
   322 
   317 subsection {* Recursion combinator for binary integers *}
       
   318 
       
   319 lemma brlem: "(bin = Int.Min) = (- bin + Int.pred 0 = 0)"
       
   320   unfolding Min_def pred_def by arith
       
   321 
       
   322 function
       
   323   bin_rec :: "'a \<Rightarrow> 'a \<Rightarrow> (int \<Rightarrow> bit \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> int \<Rightarrow> 'a"  
       
   324 where 
       
   325   "bin_rec f1 f2 f3 bin = (if bin = Int.Pls then f1 
       
   326     else if bin = Int.Min then f2
       
   327     else case bin_rl bin of (w, b) => f3 w b (bin_rec f1 f2 f3 w))"
       
   328   by pat_completeness auto
       
   329 
       
   330 termination 
       
   331   apply (relation "measure (nat o abs o snd o snd o snd)")
       
   332   apply (auto simp add: bin_rl_def bin_last_def bin_rest_def)
       
   333   unfolding Pls_def Min_def Bit0_def Bit1_def number_of_is_id
       
   334   apply auto
       
   335   done
       
   336 
       
   337 declare bin_rec.simps [simp del]
       
   338 
       
   339 lemma bin_rec_PM:
       
   340   "f = bin_rec f1 f2 f3 ==> f Int.Pls = f1 & f Int.Min = f2"
       
   341   by (auto simp add: bin_rec.simps)
       
   342 
       
   343 lemma bin_rec_Pls: "bin_rec f1 f2 f3 Int.Pls = f1"
       
   344   by (simp add: bin_rec.simps)
       
   345 
       
   346 lemma bin_rec_Min: "bin_rec f1 f2 f3 Int.Min = f2"
       
   347   by (simp add: bin_rec.simps)
       
   348 
       
   349 lemma bin_rec_Bit0:
       
   350   "f3 Int.Pls (0::bit) f1 = f1 \<Longrightarrow>
       
   351     bin_rec f1 f2 f3 (Int.Bit0 w) = f3 w (0::bit) (bin_rec f1 f2 f3 w)"
       
   352   by (simp add: bin_rec_Pls bin_rec.simps [of _ _ _ "Int.Bit0 w"])
       
   353 
       
   354 lemma bin_rec_Bit1:
       
   355   "f3 Int.Min (1::bit) f2 = f2 \<Longrightarrow>
       
   356     bin_rec f1 f2 f3 (Int.Bit1 w) = f3 w (1::bit) (bin_rec f1 f2 f3 w)"
       
   357   by (simp add: bin_rec_Min bin_rec.simps [of _ _ _ "Int.Bit1 w"])
       
   358   
       
   359 lemma bin_rec_Bit:
       
   360   "f = bin_rec f1 f2 f3  ==> f3 Int.Pls (0::bit) f1 = f1 ==> 
       
   361     f3 Int.Min (1::bit) f2 = f2 ==> f (w BIT b) = f3 w b (f w)"
       
   362   by (cases b, simp add: bin_rec_Bit0, simp add: bin_rec_Bit1)
       
   363 
       
   364 lemmas bin_rec_simps = refl [THEN bin_rec_Bit] bin_rec_Pls bin_rec_Min
       
   365   bin_rec_Bit0 bin_rec_Bit1
       
   366 
       
   367 
       
   368 subsection {* Truncating binary integers *}
   323 subsection {* Truncating binary integers *}
   369 
   324 
   370 definition
   325 definition
   371   bin_sign_def [code del] : "bin_sign = bin_rec Int.Pls Int.Min (%w b s. s)"
   326   bin_sign_def: "bin_sign k = (if k \<ge> 0 then 0 else - 1)"
   372 
   327 
   373 lemma bin_sign_simps [simp]:
   328 lemma bin_sign_simps [simp]:
   374   "bin_sign Int.Pls = Int.Pls"
   329   "bin_sign Int.Pls = Int.Pls"
   375   "bin_sign Int.Min = Int.Min"
   330   "bin_sign Int.Min = Int.Min"
   376   "bin_sign (Int.Bit0 w) = bin_sign w"
   331   "bin_sign (Int.Bit0 w) = bin_sign w"
   377   "bin_sign (Int.Bit1 w) = bin_sign w"
   332   "bin_sign (Int.Bit1 w) = bin_sign w"
   378   "bin_sign (w BIT b) = bin_sign w"
   333   "bin_sign (w BIT b) = bin_sign w"
   379   unfolding bin_sign_def by (auto simp: bin_rec_simps)
   334   by (unfold bin_sign_def numeral_simps Bit_def bitval_def) (simp_all split: bit.split)
   380 
       
   381 declare bin_sign_simps(1-4) [code]
       
   382 
   335 
   383 lemma bin_sign_rest [simp]: 
   336 lemma bin_sign_rest [simp]: 
   384   "bin_sign (bin_rest w) = (bin_sign w)"
   337   "bin_sign (bin_rest w) = bin_sign w"
   385   by (cases w rule: bin_exhaust) auto
   338   by (cases w rule: bin_exhaust) auto
   386 
   339 
   387 consts
   340 primrec bintrunc :: "nat \<Rightarrow> int \<Rightarrow> int" where
   388   bintrunc :: "nat => int => int"
       
   389 primrec 
       
   390   Z : "bintrunc 0 bin = Int.Pls"
   341   Z : "bintrunc 0 bin = Int.Pls"
   391   Suc : "bintrunc (Suc n) bin = bintrunc n (bin_rest bin) BIT (bin_last bin)"
   342 | Suc : "bintrunc (Suc n) bin = bintrunc n (bin_rest bin) BIT (bin_last bin)"
   392 
   343 
   393 consts
   344 primrec sbintrunc :: "nat => int => int" where
   394   sbintrunc :: "nat => int => int" 
       
   395 primrec 
       
   396   Z : "sbintrunc 0 bin = 
   345   Z : "sbintrunc 0 bin = 
   397     (case bin_last bin of (1::bit) => Int.Min | (0::bit) => Int.Pls)"
   346     (case bin_last bin of (1::bit) => Int.Min | (0::bit) => Int.Pls)"
   398   Suc : "sbintrunc (Suc n) bin = sbintrunc n (bin_rest bin) BIT (bin_last bin)"
   347 | Suc : "sbintrunc (Suc n) bin = sbintrunc n (bin_rest bin) BIT (bin_last bin)"
       
   348 
       
   349 lemma [code]:
       
   350   "sbintrunc 0 bin = 
       
   351     (case bin_last bin of (1::bit) => - 1 | (0::bit) => 0)"
       
   352   "sbintrunc (Suc n) bin = sbintrunc n (bin_rest bin) BIT (bin_last bin)"
       
   353   apply simp_all apply (cases "bin_last bin") apply simp apply (unfold Min_def number_of_is_id) apply simp done
   399 
   354 
   400 lemma sign_bintr:
   355 lemma sign_bintr:
   401   "!!w. bin_sign (bintrunc n w) = Int.Pls"
   356   "!!w. bin_sign (bintrunc n w) = Int.Pls"
   402   by (induct n) auto
   357   by (induct n) auto
   403 
   358 
   860 primrec bin_split :: "nat \<Rightarrow> int \<Rightarrow> int \<times> int" where
   815 primrec bin_split :: "nat \<Rightarrow> int \<Rightarrow> int \<times> int" where
   861   Z: "bin_split 0 w = (w, Int.Pls)"
   816   Z: "bin_split 0 w = (w, Int.Pls)"
   862   | Suc: "bin_split (Suc n) w = (let (w1, w2) = bin_split n (bin_rest w)
   817   | Suc: "bin_split (Suc n) w = (let (w1, w2) = bin_split n (bin_rest w)
   863         in (w1, w2 BIT bin_last w))"
   818         in (w1, w2 BIT bin_last w))"
   864 
   819 
       
   820 lemma [code]:
       
   821   "bin_split (Suc n) w = (let (w1, w2) = bin_split n (bin_rest w) in (w1, w2 BIT bin_last w))"
       
   822   "bin_split 0 w = (w, 0)"
       
   823   by (simp_all add: Pls_def)
       
   824 
   865 primrec bin_cat :: "int \<Rightarrow> nat \<Rightarrow> int \<Rightarrow> int" where
   825 primrec bin_cat :: "int \<Rightarrow> nat \<Rightarrow> int \<Rightarrow> int" where
   866   Z: "bin_cat w 0 v = w"
   826   Z: "bin_cat w 0 v = w"
   867   | Suc: "bin_cat w (Suc n) v = bin_cat w n (bin_rest v) BIT bin_last v"
   827   | Suc: "bin_cat w (Suc n) v = bin_cat w n (bin_rest v) BIT bin_last v"
   868 
   828 
   869 subsection {* Miscellaneous lemmas *}
   829 subsection {* Miscellaneous lemmas *}