src/HOL/Library/word_setup.ML
author paulson
Tue, 11 May 2004 10:47:15 +0200
changeset 14732 967db86e853c
parent 14494 48ae8d678d88
child 15013 34264f5e4691
permissions -rw-r--r--
auto update
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14494
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
     1
(*  Title:      HOL/Library/word_setup.ML
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
     2
    ID:         $Id$
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
     3
    Author:     Sebastian Skalberg (TU Muenchen)
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
     4
*)
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
     5
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
     6
local
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
     7
    fun is_const_bool (Const("True",_)) = true
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
     8
      | is_const_bool (Const("False",_)) = true
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
     9
      | is_const_bool _ = false
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    10
    fun is_const_bit (Const("Word.bit.Zero",_)) = true
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    11
      | is_const_bit (Const("Word.bit.One",_)) = true
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    12
      | is_const_bit _ = false
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    13
    fun num_is_usable (Const("Numeral.bin.Pls",_)) = true
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    14
      | num_is_usable (Const("Numeral.bin.Min",_)) = false
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    15
      | num_is_usable (Const("Numeral.bin.Bit",_) $ w $ b) =
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    16
	num_is_usable w andalso is_const_bool b
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    17
      | num_is_usable _ = false
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    18
    fun vec_is_usable (Const("List.list.Nil",_)) = true
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    19
      | vec_is_usable (Const("List.list.Cons",_) $ b $ bs) =
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    20
	vec_is_usable bs andalso is_const_bit b
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    21
      | vec_is_usable _ = false
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    22
    fun add_word thy =
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    23
	let
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    24
	    val fast1_th = PureThy.get_thm thy "Word.fast_nat_to_bv_def"
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    25
	    val fast2_th = PureThy.get_thm thy "Word.fast_bv_to_nat_def"
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    26
	    fun f sg thms (Const("Word.nat_to_bv",_) $ (Const("Numeral.number_of",_) $ t)) =
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    27
		if num_is_usable t
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    28
		then Some (Drule.cterm_instantiate [(cterm_of sg (Var(("w",0),Type("Numeral.bin",[]))),cterm_of sg t)] fast1_th)
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    29
		else None
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    30
	      | f _ _ _ = None
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    31
	    fun g sg thms (Const("Word.bv_to_nat",_) $ (t as (Const("List.list.Cons",_) $ _ $ _))) =
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    32
		if vec_is_usable t
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    33
		then Some (Drule.cterm_instantiate [(cterm_of sg (Var(("bs",0),Type("List.list",[Type("Word.bit",[])]))),cterm_of sg t)] fast2_th)
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    34
		else None
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    35
	      | g _ _ _ = None
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    36
	    val simproc1 = Simplifier.simproc (sign_of thy) "nat_to_bv" ["Word.nat_to_bv (number_of ?w)"] f
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    37
	    val simproc2 = Simplifier.simproc (sign_of thy) "bv_to_nat" ["Word.bv_to_nat (?x # ?xs)"] g
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    38
	in
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    39
	    Simplifier.change_simpset_of (op addsimprocs) [simproc1,simproc2] thy
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    40
	end
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    41
in
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    42
val setup_word = [add_word]
48ae8d678d88 Added bitvector library (Word) to HOL/Library and a theory using it (Adder)
skalberg
parents:
diff changeset
    43
end