src/HOL/SPARK/SPARK.thy
author haftmann
Mon, 02 Aug 2021 10:01:06 +0000
changeset 74101 d804e93ae9ff
parent 74097 6d7be1227d02
child 74391 930047942f46
permissions -rw-r--r--
moved theory Bit_Operations into Main corpus
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     1
(*  Title:      HOL/SPARK/SPARK.thy
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     2
    Author:     Stefan Berghofer
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     3
    Copyright:  secunet Security Networks AG
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     4
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     5
Declaration of proof functions for SPARK/Ada verification environment.
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     6
*)
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     7
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     8
theory SPARK
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
     9
imports SPARK_Setup
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    10
begin
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    11
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 54427
diff changeset
    12
text \<open>Bitwise logical operators\<close>
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    13
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    14
spark_proof_functions
74097
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 72488
diff changeset
    15
  bit__and (integer, integer) : integer = Bit_Operations.and
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 72488
diff changeset
    16
  bit__or (integer, integer) : integer = Bit_Operations.or
6d7be1227d02 organize syntax for word operations in bundles
haftmann
parents: 72488
diff changeset
    17
  bit__xor (integer, integer) : integer = Bit_Operations.xor
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    18
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    19
lemmas [simp] =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    20
  OR_upper [of _ 8, simplified zle_diff1_eq [symmetric], simplified]
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    21
  OR_upper [of _ 8, simplified]
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    22
  OR_upper [of _ 16, simplified zle_diff1_eq [symmetric], simplified]
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    23
  OR_upper [of _ 16, simplified]
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    24
  OR_upper [of _ 32, simplified zle_diff1_eq [symmetric], simplified]
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    25
  OR_upper [of _ 32, simplified]
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    26
  OR_upper [of _ 64, simplified zle_diff1_eq [symmetric], simplified]
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    27
  OR_upper [of _ 64, simplified]
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    28
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    29
lemmas [simp] =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    30
  XOR_upper [of _ 8, simplified zle_diff1_eq [symmetric], simplified]
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    31
  XOR_upper [of _ 8, simplified]
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    32
  XOR_upper [of _ 16, simplified zle_diff1_eq [symmetric], simplified]
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    33
  XOR_upper [of _ 16, simplified]
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    34
  XOR_upper [of _ 32, simplified zle_diff1_eq [symmetric], simplified]
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    35
  XOR_upper [of _ 32, simplified]
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    36
  XOR_upper [of _ 64, simplified zle_diff1_eq [symmetric], simplified]
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    37
  XOR_upper [of _ 64, simplified]
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    38
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    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: 70185
diff changeset
    40
  "NOT (word_of_int x :: ('a::len) word) =
70185
ac1706cdde25 clarified notation
haftmann
parents: 67399
diff changeset
    41
  word_of_int (2 ^ LENGTH('a) - 1 - x)"
72488
ee659bca8955 factored out theory Bits_Int
haftmann
parents: 71954
diff changeset
    42
  by (simp flip: mask_eq_exp_minus_1 add: of_int_mask_eq not_eq_complement)
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    43
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    44
lemmas [simp] =
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    45
  bit_not_spark_eq [where 'a=8, simplified]
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    46
  bit_not_spark_eq [where 'a=16, simplified]
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    47
  bit_not_spark_eq [where 'a=32, simplified]
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    48
  bit_not_spark_eq [where 'a=64, simplified]
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    49
50788
fec14e8fefb8 Added proof function declarations for min and max
berghofe
parents: 47108
diff changeset
    50
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 54427
diff changeset
    51
text \<open>Minimum and maximum\<close>
50788
fec14e8fefb8 Added proof function declarations for min and max
berghofe
parents: 47108
diff changeset
    52
fec14e8fefb8 Added proof function declarations for min and max
berghofe
parents: 47108
diff changeset
    53
spark_proof_functions
fec14e8fefb8 Added proof function declarations for min and max
berghofe
parents: 47108
diff changeset
    54
  integer__min = "min :: int \<Rightarrow> int \<Rightarrow> int"
fec14e8fefb8 Added proof function declarations for min and max
berghofe
parents: 47108
diff changeset
    55
  integer__max = "max :: int \<Rightarrow> int \<Rightarrow> int"
fec14e8fefb8 Added proof function declarations for min and max
berghofe
parents: 47108
diff changeset
    56
41561
d1318f3c86ba Added new SPARK verification environment.
berghofe
parents:
diff changeset
    57
end
54427
783861a66a60 separated comparision on bit operations into separate theory
haftmann
parents: 50788
diff changeset
    58