src/HOL/SMT_Examples/SMT_Word_Examples.thy
author wenzelm
Sat, 16 Aug 2014 12:15:56 +0200
changeset 57947 189d421ca72d
parent 56109 1ba56358eba4
child 58061 3d060f43accb
permissions -rw-r--r--
updated syntax for localized commands;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
     1
(*  Title:      HOL/SMT_Examples/SMT_Word_Examples.thy
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
     2
    Author:     Sascha Boehme, TU Muenchen
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
     3
*)
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
     4
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
     5
header {* Word examples for for SMT binding *}
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
     6
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
     7
theory SMT_Word_Examples
56046
683148f3ae48 full path
blanchet
parents: 47152
diff changeset
     8
imports "~~/src/HOL/Word/Word"
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
     9
begin
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    10
56079
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    11
declare [[smt2_oracle = true]]
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    12
declare [[z3_new_extensions = true]]
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    13
declare [[smt2_certificates = "SMT_Word_Examples.certs2"]]
56109
1ba56358eba4 updated SMT example certificates
blanchet
parents: 56079
diff changeset
    14
declare [[smt2_read_only_certificates = true]]
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    15
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    16
text {*
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    17
Currently, there is no proof reconstruction for words.
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    18
All lemmas are proved using the oracle mechanism.
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    19
*}
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    20
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    21
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    22
section {* Bitvector numbers *}
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    23
56079
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    24
lemma "(27 :: 4 word) = -5" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    25
lemma "(27 :: 4 word) = 11" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    26
lemma "23 < (27::8 word)" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    27
lemma "27 + 11 = (6::5 word)" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    28
lemma "7 * 3 = (21::8 word)" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    29
lemma "11 - 27 = (-16::8 word)" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    30
lemma "- -11 = (11::5 word)" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    31
lemma "-40 + 1 = (-39::7 word)" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    32
lemma "a + 2 * b + c - b = (b + c) + (a :: 32 word)" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    33
lemma "x = (5 :: 4 word) \<Longrightarrow> 4 * x = 4" by smt2
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    34
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    35
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    36
section {* Bit-level logic *}
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    37
56079
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    38
lemma "0b110 AND 0b101 = (0b100 :: 32 word)" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    39
lemma "0b110 OR 0b011 = (0b111 :: 8 word)" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    40
lemma "0xF0 XOR 0xFF = (0x0F :: 8 word)" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    41
lemma "NOT (0xF0 :: 16 word) = 0xFF0F" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    42
lemma "word_cat (27::4 word) (27::8 word) = (2843::12 word)" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    43
lemma "word_cat (0b0011::4 word) (0b1111::6word) = (0b0011001111 :: 10 word)" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    44
lemma "slice 1 (0b10110 :: 4 word) = (0b11 :: 2 word)" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    45
lemma "ucast (0b1010 :: 4 word) = (0b1010 :: 10 word)" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    46
lemma "scast (0b1010 :: 4 word) = (0b111010 :: 6 word)" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    47
lemma "0b10011 << 2 = (0b1001100::8 word)" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    48
lemma "0b11001 >> 2 = (0b110::8 word)" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    49
lemma "0b10011 >>> 2 = (0b100::8 word)" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    50
lemma "word_rotr 2 0b0110 = (0b1001::4 word)" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    51
lemma "word_rotl 1 0b1110 = (0b1101::4 word)" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    52
lemma "(x AND 0xff00) OR (x AND 0x00ff) = (x::16 word)" by smt2
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    53
lemma "w < 256 \<Longrightarrow> (w :: 16 word) AND 0x00FF = w" by smt2
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    54
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    55
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    56
section {* Combined integer-bitvector properties *}
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    57
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    58
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    59
  assumes "bv2int 0 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    60
      and "bv2int 1 = 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    61
      and "bv2int 2 = 2"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    62
      and "bv2int 3 = 3"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    63
      and "\<forall>x::2 word. bv2int x > 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    64
  shows "\<forall>i::int. i < 0 \<longrightarrow> (\<forall>x::2 word. bv2int x > i)"
56079
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    65
  using assms by smt2
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    66
56079
175ac95720d4 use 'smt2' in SMT examples as much as currently possible
blanchet
parents: 56046
diff changeset
    67
lemma "P (0 \<le> (a :: 4 word)) = P True" by smt2
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    68
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    69
end