src/HOL/Library/Code_Target_Numeral_Float.thy
author paulson <lp15@cam.ac.uk>
Sun, 03 Aug 2025 20:34:24 +0100
changeset 82913 7c870287f04f
parent 73546 7cb3fefef79e
permissions -rw-r--r--
New lemmas about improper integrals and other things
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
73546
7cb3fefef79e confluent preprocessing for floats in presence of target language numerals
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Library/Code_Target_Numeral_Float.thy
7cb3fefef79e confluent preprocessing for floats in presence of target language numerals
haftmann
parents:
diff changeset
     2
    Author:     Florian Haftmann, TU Muenchen
7cb3fefef79e confluent preprocessing for floats in presence of target language numerals
haftmann
parents:
diff changeset
     3
*)
7cb3fefef79e confluent preprocessing for floats in presence of target language numerals
haftmann
parents:
diff changeset
     4
7cb3fefef79e confluent preprocessing for floats in presence of target language numerals
haftmann
parents:
diff changeset
     5
section \<open>Preprocessor setup for floats implemented by target language numerals\<close>
7cb3fefef79e confluent preprocessing for floats in presence of target language numerals
haftmann
parents:
diff changeset
     6
7cb3fefef79e confluent preprocessing for floats in presence of target language numerals
haftmann
parents:
diff changeset
     7
theory Code_Target_Numeral_Float
7cb3fefef79e confluent preprocessing for floats in presence of target language numerals
haftmann
parents:
diff changeset
     8
imports Float Code_Target_Numeral
7cb3fefef79e confluent preprocessing for floats in presence of target language numerals
haftmann
parents:
diff changeset
     9
begin
7cb3fefef79e confluent preprocessing for floats in presence of target language numerals
haftmann
parents:
diff changeset
    10
7cb3fefef79e confluent preprocessing for floats in presence of target language numerals
haftmann
parents:
diff changeset
    11
lemma numeral_float_computation_unfold [code_computation_unfold]:
7cb3fefef79e confluent preprocessing for floats in presence of target language numerals
haftmann
parents:
diff changeset
    12
  \<open>numeral k = Float (int_of_integer (Code_Numeral.positive k)) 0\<close>
7cb3fefef79e confluent preprocessing for floats in presence of target language numerals
haftmann
parents:
diff changeset
    13
  \<open>- numeral k = Float (int_of_integer (Code_Numeral.negative k)) 0\<close>
7cb3fefef79e confluent preprocessing for floats in presence of target language numerals
haftmann
parents:
diff changeset
    14
  by (simp_all add: Float.compute_float_numeral Float.compute_float_neg_numeral)
7cb3fefef79e confluent preprocessing for floats in presence of target language numerals
haftmann
parents:
diff changeset
    15
7cb3fefef79e confluent preprocessing for floats in presence of target language numerals
haftmann
parents:
diff changeset
    16
end