src/HOL/SMT/SMT_Base.thy
author boehmes
Wed, 12 May 2010 23:53:55 +0200
changeset 36891 e0d295cb8bfd
parent 36884 88cf4896b980
child 36893 48cf03469dc6
permissions -rw-r--r--
move the addition of extra facts into a separate module
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33010
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
     1
(*  Title:      HOL/SMT/SMT_Base.thy
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
     2
    Author:     Sascha Boehme, TU Muenchen
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
     3
*)
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
     4
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
     5
header {* SMT-specific definitions and basic tools *}
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
     6
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
     7
theory SMT_Base
35105
1822c658a5e4 use full paths when importing theories
boehmes
parents: 34960
diff changeset
     8
imports Real "~~/src/HOL/Word/Word"
33010
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
     9
uses
35943
51b9155467cc cache_io is now just a single ML file instead of a component
boehmes
parents: 35151
diff changeset
    10
  "~~/src/Tools/cache_io.ML"
36891
e0d295cb8bfd move the addition of extra facts into a separate module
boehmes
parents: 36884
diff changeset
    11
  ("Tools/smt_additional_facts.ML")
33010
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    12
  ("Tools/smt_normalize.ML")
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    13
  ("Tools/smt_monomorph.ML")
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    14
  ("Tools/smt_translate.ML")
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    15
  ("Tools/smt_solver.ML")
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    16
  ("Tools/smtlib_interface.ML")
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    17
begin
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    18
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    19
section {* Triggers for quantifier instantiation *}
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    20
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    21
text {*
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    22
Some SMT solvers support triggers for quantifier instantiation. Each trigger
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    23
consists of one ore more patterns. A pattern may either be a list of positive
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    24
subterms (the first being tagged by "pat" and the consecutive subterms tagged
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    25
by "andpat"), or a list of negative subterms (the first being tagged by "nopat"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    26
and the consecutive subterms tagged by "andpat").
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    27
*}
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    28
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    29
datatype pattern = Pattern
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    30
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    31
definition pat :: "'a \<Rightarrow> pattern"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    32
where "pat _ = Pattern"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    33
33249
2b65e9ed2e6e removed unused file smt_builtin.ML,
boehmes
parents: 33010
diff changeset
    34
definition nopat :: "'a \<Rightarrow> pattern"
33010
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    35
where "nopat _ = Pattern"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    36
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    37
definition andpat :: "pattern \<Rightarrow> 'a \<Rightarrow> pattern" (infixl "andpat" 60)
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    38
where "_ andpat _ = Pattern"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    39
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    40
definition trigger :: "pattern list \<Rightarrow> bool \<Rightarrow> bool"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    41
where "trigger _ P = P"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    42
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    43
36884
88cf4896b980 moved the addition of DLO tactic into the Z3 theory (DLO is required only for Z3 proof reconstruction)
boehmes
parents: 36084
diff changeset
    44
33010
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    45
section {* Arithmetic *}
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    46
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    47
text {*
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    48
The sign of @{term "op mod :: int \<Rightarrow> int \<Rightarrow> int"} follows the sign of the
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    49
divisor. In contrast to that, the sign of the following operation is that of
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    50
the dividend.
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    51
*}
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    52
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    53
definition rem :: "int \<Rightarrow> int \<Rightarrow> int" (infixl "rem" 70)
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    54
where "a rem b = 
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    55
  (if (a \<ge> 0 \<and> b < 0) \<or> (a < 0 \<and> b \<ge> 0) then - (a mod b) else a mod b)"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    56
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    57
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    58
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    59
section {* Bitvectors *}
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    60
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    61
text {*
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    62
The following definitions provide additional functions not found in HOL-Word.
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    63
*}
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    64
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    65
definition sdiv :: "'a::len word \<Rightarrow> 'a word \<Rightarrow> 'a word" (infix "sdiv" 70)
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    66
where "w1 sdiv w2 = word_of_int (sint w1 div sint w2)"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    67
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    68
definition smod :: "'a::len word \<Rightarrow> 'a word \<Rightarrow> 'a word" (infix "smod" 70)
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    69
  (* sign follows divisor *)
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    70
where "w1 smod w2 = word_of_int (sint w1 mod sint w2)"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    71
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    72
definition srem :: "'a::len word \<Rightarrow> 'a word \<Rightarrow> 'a word" (infix "srem" 70)
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    73
  (* sign follows dividend *)
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    74
where "w1 srem w2 = word_of_int (sint w1 rem sint w2)"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    75
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    76
definition bv_shl :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> 'a word"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    77
where "bv_shl w1 w2 = (w1 << unat w2)"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    78
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    79
definition bv_lshr :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> 'a word"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    80
where "bv_lshr w1 w2 = (w1 >> unat w2)"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    81
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    82
definition bv_ashr :: "'a::len word \<Rightarrow> 'a word \<Rightarrow> 'a word"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    83
where "bv_ashr w1 w2 = (w1 >>> unat w2)"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    84
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    85
36884
88cf4896b980 moved the addition of DLO tactic into the Z3 theory (DLO is required only for Z3 proof reconstruction)
boehmes
parents: 36084
diff changeset
    86
33010
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    87
section {* Higher-Order Encoding *}
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    88
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    89
definition "apply" where "apply f x = f x"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
    90
34960
1d5ee19ef940 support skolem constant for extensional arrays in Z3 proofs
boehmes
parents: 33249
diff changeset
    91
definition array_ext where "array_ext a b = (SOME x. a = b \<or> a x \<noteq> b x)"
1d5ee19ef940 support skolem constant for extensional arrays in Z3 proofs
boehmes
parents: 33249
diff changeset
    92
1d5ee19ef940 support skolem constant for extensional arrays in Z3 proofs
boehmes
parents: 33249
diff changeset
    93
lemma fun_upd_eq: "(f = f (x := y)) = (f x = y)"
1d5ee19ef940 support skolem constant for extensional arrays in Z3 proofs
boehmes
parents: 33249
diff changeset
    94
proof
1d5ee19ef940 support skolem constant for extensional arrays in Z3 proofs
boehmes
parents: 33249
diff changeset
    95
  assume "f = f(x:=y)"
1d5ee19ef940 support skolem constant for extensional arrays in Z3 proofs
boehmes
parents: 33249
diff changeset
    96
  hence "f x = (f(x:=y)) x" by simp
1d5ee19ef940 support skolem constant for extensional arrays in Z3 proofs
boehmes
parents: 33249
diff changeset
    97
  thus "f x = y" by simp
1d5ee19ef940 support skolem constant for extensional arrays in Z3 proofs
boehmes
parents: 33249
diff changeset
    98
qed (auto simp add: ext)
1d5ee19ef940 support skolem constant for extensional arrays in Z3 proofs
boehmes
parents: 33249
diff changeset
    99
1d5ee19ef940 support skolem constant for extensional arrays in Z3 proofs
boehmes
parents: 33249
diff changeset
   100
lemmas array_rules =
1d5ee19ef940 support skolem constant for extensional arrays in Z3 proofs
boehmes
parents: 33249
diff changeset
   101
  ext fun_upd_apply fun_upd_same fun_upd_other fun_upd_upd fun_upd_eq apply_def
33010
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   102
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   103
36884
88cf4896b980 moved the addition of DLO tactic into the Z3 theory (DLO is required only for Z3 proof reconstruction)
boehmes
parents: 36084
diff changeset
   104
33010
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   105
section {* First-order logic *}
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   106
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   107
text {*
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   108
Some SMT solver formats require a strict separation between formulas and terms.
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   109
The following marker symbols are used internally to separate those categories:
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   110
*}
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   111
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   112
definition formula :: "bool \<Rightarrow> bool" where "formula x = x"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   113
definition "term" where "term x = x"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   114
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   115
text {*
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   116
Predicate symbols also occurring as function symbols are turned into function
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   117
symbols by translating atomic formulas into terms:
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   118
*}
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   119
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   120
abbreviation holds :: "bool \<Rightarrow> bool" where "holds \<equiv> (\<lambda>P. term P = term True)"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   121
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   122
text {*
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   123
The following constant represents equivalence, to be treated differently than
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   124
the (polymorphic) equality predicate:
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   125
*}
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   126
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   127
definition iff :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infix "iff" 50) where
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   128
  "(x iff y) = (x = y)"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   129
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   130
36884
88cf4896b980 moved the addition of DLO tactic into the Z3 theory (DLO is required only for Z3 proof reconstruction)
boehmes
parents: 36084
diff changeset
   131
33010
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   132
section {* Setup *}
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   133
36891
e0d295cb8bfd move the addition of extra facts into a separate module
boehmes
parents: 36884
diff changeset
   134
use "Tools/smt_additional_facts.ML"
33010
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   135
use "Tools/smt_normalize.ML"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   136
use "Tools/smt_monomorph.ML"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   137
use "Tools/smt_translate.ML"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   138
use "Tools/smt_solver.ML"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   139
use "Tools/smtlib_interface.ML"
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   140
36084
3176ec2244ad always unfold definitions of specific constants (including special binders)
boehmes
parents: 35943
diff changeset
   141
setup {* SMT_Solver.setup *}
33010
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   142
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents:
diff changeset
   143
end