| author | wenzelm | 
| Tue, 26 Jan 2021 23:34:40 +0100 | |
| changeset 73194 | c0d6d57a9a31 | 
| parent 72488 | ee659bca8955 | 
| child 74097 | 6d7be1227d02 | 
| permissions | -rw-r--r-- | 
| 41561 | 1 | (* Title: HOL/SPARK/SPARK.thy | 
| 2 | Author: Stefan Berghofer | |
| 3 | Copyright: secunet Security Networks AG | |
| 4 | ||
| 5 | Declaration of proof functions for SPARK/Ada verification environment. | |
| 6 | *) | |
| 7 | ||
| 8 | theory SPARK | |
| 9 | imports SPARK_Setup | |
| 10 | begin | |
| 11 | ||
| 63167 | 12 | text \<open>Bitwise logical operators\<close> | 
| 41561 | 13 | |
| 14 | spark_proof_functions | |
| 67399 | 15 | bit__and (integer, integer) : integer = "(AND)" | 
| 16 | bit__or (integer, integer) : integer = "(OR)" | |
| 17 | bit__xor (integer, integer) : integer = "(XOR)" | |
| 41561 | 18 | |
| 19 | lemmas [simp] = | |
| 20 | OR_upper [of _ 8, simplified zle_diff1_eq [symmetric], simplified] | |
| 21 | OR_upper [of _ 8, simplified] | |
| 22 | OR_upper [of _ 16, simplified zle_diff1_eq [symmetric], simplified] | |
| 23 | OR_upper [of _ 16, simplified] | |
| 24 | OR_upper [of _ 32, simplified zle_diff1_eq [symmetric], simplified] | |
| 25 | OR_upper [of _ 32, simplified] | |
| 26 | OR_upper [of _ 64, simplified zle_diff1_eq [symmetric], simplified] | |
| 27 | OR_upper [of _ 64, simplified] | |
| 28 | ||
| 29 | lemmas [simp] = | |
| 30 | XOR_upper [of _ 8, simplified zle_diff1_eq [symmetric], simplified] | |
| 31 | XOR_upper [of _ 8, simplified] | |
| 32 | XOR_upper [of _ 16, simplified zle_diff1_eq [symmetric], simplified] | |
| 33 | XOR_upper [of _ 16, simplified] | |
| 34 | XOR_upper [of _ 32, simplified zle_diff1_eq [symmetric], simplified] | |
| 35 | XOR_upper [of _ 32, simplified] | |
| 36 | XOR_upper [of _ 64, simplified zle_diff1_eq [symmetric], simplified] | |
| 37 | XOR_upper [of _ 64, simplified] | |
| 38 | ||
| 39 | lemma bit_not_spark_eq: | |
| 71954 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 haftmann parents: 
70185diff
changeset | 40 |   "NOT (word_of_int x :: ('a::len) word) =
 | 
| 70185 | 41 |   word_of_int (2 ^ LENGTH('a) - 1 - x)"
 | 
| 72488 | 42 | by (simp flip: mask_eq_exp_minus_1 add: of_int_mask_eq not_eq_complement) | 
| 41561 | 43 | |
| 44 | lemmas [simp] = | |
| 45 | bit_not_spark_eq [where 'a=8, simplified] | |
| 46 | bit_not_spark_eq [where 'a=16, simplified] | |
| 47 | bit_not_spark_eq [where 'a=32, simplified] | |
| 48 | bit_not_spark_eq [where 'a=64, simplified] | |
| 49 | ||
| 50788 
fec14e8fefb8
Added proof function declarations for min and max
 berghofe parents: 
47108diff
changeset | 50 | |
| 63167 | 51 | text \<open>Minimum and maximum\<close> | 
| 50788 
fec14e8fefb8
Added proof function declarations for min and max
 berghofe parents: 
47108diff
changeset | 52 | |
| 
fec14e8fefb8
Added proof function declarations for min and max
 berghofe parents: 
47108diff
changeset | 53 | spark_proof_functions | 
| 
fec14e8fefb8
Added proof function declarations for min and max
 berghofe parents: 
47108diff
changeset | 54 | integer__min = "min :: int \<Rightarrow> int \<Rightarrow> int" | 
| 
fec14e8fefb8
Added proof function declarations for min and max
 berghofe parents: 
47108diff
changeset | 55 | integer__max = "max :: int \<Rightarrow> int \<Rightarrow> int" | 
| 
fec14e8fefb8
Added proof function declarations for min and max
 berghofe parents: 
47108diff
changeset | 56 | |
| 41561 | 57 | end | 
| 54427 
783861a66a60
separated comparision on bit operations into separate theory
 haftmann parents: 
50788diff
changeset | 58 |