Theory IEEE
section ‹Specification of the IEEE standard›
theory IEEE
imports
"HOL-Library.Float"
Word_Lib.Word_Lemmas
begin
typedef (overloaded) ('e::len, 'f::len) float = "UNIV::(1 word × 'e word × 'f word) set"
by auto
setup_lifting type_definition_float
syntax "_float" :: "type ⇒ type ⇒ type" (‹'(_, _') float›)
syntax_types "_float" ⇌ float
text ‹parse ‹('a, 'b) float› as ('a::len, 'b::len) float.›
parse_translation ‹
let
fun float t u = Syntax.const @{type_syntax float} $ t $ u;
fun len_tr u =
(case Term_Position.strip_positions u of
v as Free (x, _) =>
if Lexicon.is_tid x then
(Syntax.const @{syntax_const "_ofsort"} $ v $
Syntax.const @{class_syntax len})
else u
| _ => u)
fun len_float_tr [t, u] =
float (len_tr t) (len_tr u)
in
[(@{syntax_const "_float"}, K len_float_tr)]
end
›
subsection ‹Derived parameters for floating point formats›
definition wordlength :: "('e, 'f) float itself ⇒ nat"
where "wordlength x = LENGTH('e) + LENGTH('f) + 1"
definition bias :: "('e, 'f) float itself ⇒ nat"
where "bias x = 2^(LENGTH('e) - 1) - 1"
definition emax :: "('e, 'f) float itself ⇒ nat"
where "emax x = unat (- 1::'e word)"
abbreviation fracwidth::"('e, 'f) float itself ⇒ nat" where
"fracwidth _ ≡ LENGTH('f)"
subsection ‹Predicates for the four IEEE formats›
definition is_single :: "('e, 'f) float itself ⇒ bool"
where "is_single x ⟷ LENGTH('e) = 8 ∧ wordlength x = 32"
definition is_double :: "('e, 'f) float itself ⇒ bool"
where "is_double x ⟷ LENGTH('e) = 11 ∧ wordlength x = 64"
definition is_single_extended :: "('e, 'f) float itself ⇒ bool"
where "is_single_extended x ⟷ LENGTH('e) ≥ 11 ∧ wordlength x ≥ 43"
definition is_double_extended :: "('e, 'f) float itself ⇒ bool"
where "is_double_extended x ⟷ LENGTH('e) ≥ 15 ∧ wordlength x ≥ 79"
subsection ‹Extractors for fields›
lift_definition sign::"('e, 'f) float ⇒ nat" is
"λ(s::1 word, _::'e word, _::'f word). unat s" .
lift_definition exponent::"('e, 'f) float ⇒ nat" is
"λ(_, e::'e word, _). unat e" .
lift_definition fraction::"('e, 'f) float ⇒ nat" is
"λ(_, _, f::'f word). unat f" .
abbreviation "real_of_word x ≡ real (unat x)"
lift_definition valof :: "('e, 'f) float ⇒ real"
is "λ(s, e, f).
let x = (TYPE(('e, 'f) float)) in
(if e = 0
then (-1::real)^(unat s) * (2 / (2^bias x)) * (real_of_word f/2^(LENGTH('f)))
else (-1::real)^(unat s) * ((2^(unat e)) / (2^bias x)) * (1 + real_of_word f/2^LENGTH('f)))"
.
subsection ‹Partition of numbers into disjoint classes›
definition is_nan :: "('e, 'f) float ⇒ bool"
where "is_nan a ⟷ exponent a = emax TYPE(('e, 'f)float) ∧ fraction a ≠ 0"
definition is_infinity :: "('e, 'f) float ⇒ bool"
where "is_infinity a ⟷ exponent a = emax TYPE(('e, 'f)float) ∧ fraction a = 0"
definition is_normal :: "('e, 'f) float ⇒ bool"
where "is_normal a ⟷ 0 < exponent a ∧ exponent a < emax TYPE(('e, 'f)float)"
definition is_denormal :: "('e, 'f) float ⇒ bool"
where "is_denormal a ⟷ exponent a = 0 ∧ fraction a ≠ 0"
definition is_zero :: "('e, 'f) float ⇒ bool"
where "is_zero a ⟷ exponent a = 0 ∧ fraction a = 0"
definition is_finite :: "('e, 'f) float ⇒ bool"
where "is_finite a ⟷ (is_normal a ∨ is_denormal a ∨ is_zero a)"
subsection ‹Special values›
lift_definition plus_infinity :: "('e, 'f) float" (‹∞›) is "(0, - 1, 0)" .
lift_definition topfloat :: "('e, 'f) float" is "(0, - 2, 2^LENGTH('f) - 1)" .
instantiation float::(len, len) zero begin
lift_definition zero_float :: "('e, 'f) float" is "(0, 0, 0)" .
instance proof qed
end
subsection ‹Negation operation on floating point values›
instantiation float::(len, len) uminus begin
lift_definition uminus_float :: "('e, 'f) float ⇒ ('e, 'f) float" is "λ(s, e, f). (1 - s, e, f)" .
instance proof qed
end
abbreviation (input) "minus_zero ≡ - (0::('e, 'f)float)"
abbreviation (input) "minus_infinity ≡ - ∞"
abbreviation (input) "bottomfloat ≡ - topfloat"
subsection ‹Real number valuations›
text ‹The largest value that can be represented in floating point format.›
definition largest :: "('e, 'f) float itself ⇒ real"
where "largest x = (2^(emax x - 1) / 2^bias x) * (2 - 1/(2^fracwidth x))"
text ‹Threshold, used for checking overflow.›
definition threshold :: "('e, 'f) float itself ⇒ real"
where "threshold x = (2^(emax x - 1) / 2^bias x) * (2 - 1/(2^(Suc(fracwidth x))))"
text ‹Unit of least precision.›
lift_definition one_lp::"('e ,'f) float ⇒ ('e ,'f) float" is "λ(s, e, f). (0, e::'e word, 1)" .
lift_definition zero_lp::"('e ,'f) float ⇒ ('e ,'f) float" is "λ(s, e, f). (0, e::'e word, 0)" .
definition ulp :: "('e, 'f) float ⇒ real" where "ulp a = valof (one_lp a) - valof (zero_lp a)"
text ‹Enumerated type for rounding modes.›