Trying to generalize a bit vector that uses typedef, bool list, and nat length

僤鯓⒐⒋嵵緔 提交于 2019-12-06 11:08:52

Isabelle does not support dependent types, but there are ways to still do what you want to do. For instance, there is already a stack of type classes and type syntax for type-level natural numbers.

theory Scratch
imports Main "~~/src/HOL/Library/Numeral_Type"
begin

lemma "(UNIV :: 4 set) = {0,1,2,3}"
  by (subst UNIV_enum) eval

As you can see, the type 4 is a type that contains the numbers from 0 to 3. Incidentally, this can also be used for computations in modular arithmetic:

lemma "((2 + 3) :: 4) = 1" by simp
lemma "((2 * 3) :: 4) = 2" by simp

You can use these numeral types to parametrise your bit vectors with a length:

typedef ('n::finite) bitvec = "{bs :: bool list. length bs = CARD('n)}"
  morphisms bitvec_to_list Abs_bitvec
  by (simp add: Ex_list_of_length)

setup_lifting type_definition_bitvec

You can access the n-th element of a bit vector by lifting the nth function from Boolean lists to bit vectors, which works automatically:

lift_definition nth :: "('n::finite) bitvec ⇒ nat ⇒ bool" (infixl "$" 90) is List.nth .

Converting boolean lists to bit vectors is a bit tricky, because the list you get in might not have the correct length; the expression list_to_bitvec [True] :: 2 bitvec would typecheck, but is obviously problematic. You could solve this either by returning undefined or, perhaps more appropriate in this instance, filling up the list with False or truncating it to get the right length:

definition set_length :: "nat ⇒ bool list ⇒ bool list" where
  "set_length n xs = (if length xs < n then xs @ replicate (n - length xs) False else take n xs)"

lemma length_set_length[simp]: "length (set_length n xs) = n"
  unfolding set_length_def by auto

Now we can define a function that converts a list of Booleans to a bit vector:

definition list_to_bitvec :: "bool list ⇒ ('n::finite) bitvec" where
  "list_to_bitvec xs = Abs_bitvec (set_length CARD('n) xs)"

However, we are not allowed to use Abs_bitvec in code equations; if you tried to evaluate, say, list_to_bitvec [True] :: 1 bitvec, you would get an abstraction violation. We have to give an explicit code abstract equation in terms of the morphism list_to_bitvec:

lemma list_to_bitvec_code[code abstract]:
  "bitvec_to_list (list_to_bitvec xs :: ('n::finite) bitvec) = set_length CARD('n) xs"
  unfolding list_to_bitvec_def by (simp add: Abs_bitvec_inverse)

And now we are basically done and can do e.g. this:

definition myvec :: "4 bitvec" where "myvec = list_to_bitvec [True, False, True]"

value myvec
(* Output: "Abs_bitvec [True, False, True, False]" :: "4 bitvec" *)

value "myvec $ 2"
(* Output: "True" :: "bool" *)

Note that you always have to annotate the result of list_to_bitvec with its length; Isabelle can not infer the length.

You may also want to have a look at the Word theory in ~~/src/HOL/Word/; it implements machine words of fixed length with all kinds of bit operations like NOT, AND, OR, etc.:

value "42 AND 23 :: 32 word"
(* Output: "2" :: "32 word" *)

value "293 :: 8 word"
(* Output: "37" :: "8 word" *)

value "test_bit (42 :: 8 word) 1"
(* Output: "True" :: "bool" *)

value "set_bit (42 :: 8 word) 2 True"
(* Output: "46" :: "8 word" *)

value "(BITS i. i < 4) :: 8 word"
(* Output: "15" :: "8 word" *)

Another related type are the vectors in src/HOL/Library/Multivariate_Analysis/Finite_Cartesian_Product.

标签
易学教程内所有资源均来自网络或用户发布的内容,如有违反法律规定的内容欢迎反馈
该文章没有解决你所遇到的问题?点击提问,说说你的问题,让更多的人一起探讨吧!