src/HOL/SMT.thy
author wenzelm
Wed, 03 Nov 2010 21:53:56 +0100
changeset 40335 3e4bb6e7c3ca
parent 40277 4e3a3461c1a6
child 40424 7550b2cba1cb
permissions -rw-r--r--
feeder: treat header as escaped utf8 to allow initial ML text to refer to non-ASCII file/directory names (e.g. "Documents/" on Chinese Ubuntu);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     1
(*  Title:      HOL/SMT.thy
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     2
    Author:     Sascha Boehme, TU Muenchen
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     3
*)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     4
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     5
header {* Bindings to Satisfiability Modulo Theories (SMT) solvers *}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     6
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     7
theory SMT
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     8
imports List
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     9
uses
39483
9f0e5684f04b add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
boehmes
parents: 39298
diff changeset
    10
  "Tools/Datatype/datatype_selectors.ML"
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    11
  ("Tools/SMT/smt_monomorph.ML")
40277
4e3a3461c1a6 added crafted list of SMT built-in constants
boehmes
parents: 40274
diff changeset
    12
  ("Tools/SMT/smt_builtin.ML")
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    13
  ("Tools/SMT/smt_normalize.ML")
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    14
  ("Tools/SMT/smt_translate.ML")
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    15
  ("Tools/SMT/smt_solver.ML")
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    16
  ("Tools/SMT/smtlib_interface.ML")
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    17
  ("Tools/SMT/z3_proof_parser.ML")
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    18
  ("Tools/SMT/z3_proof_tools.ML")
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    19
  ("Tools/SMT/z3_proof_literals.ML")
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    20
  ("Tools/SMT/z3_proof_reconstruction.ML")
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    21
  ("Tools/SMT/z3_model.ML")
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    22
  ("Tools/SMT/z3_interface.ML")
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
    23
  ("Tools/SMT/smt_setup_solvers.ML")
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    24
begin
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    25
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    26
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    27
36902
c6bae4456741 use 'subsection' instead of 'section', to maintain 1 chapter per file in generated document
huffman
parents: 36899
diff changeset
    28
subsection {* Triggers for quantifier instantiation *}
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    29
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    30
text {*
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    31
Some SMT solvers support triggers for quantifier instantiation.
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    32
Each trigger consists of one ore more patterns.  A pattern may either
37124
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36902
diff changeset
    33
be a list of positive subterms (each being tagged by "pat"), or a
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36902
diff changeset
    34
list of negative subterms (each being tagged by "nopat").
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36902
diff changeset
    35
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36902
diff changeset
    36
When an SMT solver finds a term matching a positive pattern (a
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36902
diff changeset
    37
pattern with positive subterms only), it instantiates the
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36902
diff changeset
    38
corresponding quantifier accordingly.  Negative patterns inhibit
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36902
diff changeset
    39
quantifier instantiations.  Each pattern should mention all preceding
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36902
diff changeset
    40
bound variables.
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    41
*}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    42
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    43
datatype pattern = Pattern
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    44
37124
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36902
diff changeset
    45
definition pat :: "'a \<Rightarrow> pattern" where "pat _ = Pattern"
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36902
diff changeset
    46
definition nopat :: "'a \<Rightarrow> pattern" where "nopat _ = Pattern"
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    47
37124
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36902
diff changeset
    48
definition trigger :: "pattern list list \<Rightarrow> bool \<Rightarrow> bool"
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    49
where "trigger _ P = P"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    50
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    51
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    52
40274
6486c610a549 introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
boehmes
parents: 40162
diff changeset
    53
subsection {* Distinctness *}
6486c610a549 introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
boehmes
parents: 40162
diff changeset
    54
6486c610a549 introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
boehmes
parents: 40162
diff changeset
    55
text {*
6486c610a549 introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
boehmes
parents: 40162
diff changeset
    56
As an abbreviation for a quadratic number of inequalities, SMT solvers
6486c610a549 introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
boehmes
parents: 40162
diff changeset
    57
provide a built-in @{text distinct}.  To avoid confusion with the
6486c610a549 introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
boehmes
parents: 40162
diff changeset
    58
already defined (and more general) @{term List.distinct}, a separate
6486c610a549 introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
boehmes
parents: 40162
diff changeset
    59
constant is defined.
6486c610a549 introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
boehmes
parents: 40162
diff changeset
    60
*}
6486c610a549 introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
boehmes
parents: 40162
diff changeset
    61
6486c610a549 introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
boehmes
parents: 40162
diff changeset
    62
definition distinct :: "'a list \<Rightarrow> bool"
6486c610a549 introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
boehmes
parents: 40162
diff changeset
    63
where "distinct xs = List.distinct xs"
6486c610a549 introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
boehmes
parents: 40162
diff changeset
    64
6486c610a549 introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
boehmes
parents: 40162
diff changeset
    65
6486c610a549 introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
boehmes
parents: 40162
diff changeset
    66
36902
c6bae4456741 use 'subsection' instead of 'section', to maintain 1 chapter per file in generated document
huffman
parents: 36899
diff changeset
    67
subsection {* Higher-order encoding *}
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    68
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    69
text {*
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    70
Application is made explicit for constants occurring with varying
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    71
numbers of arguments.  This is achieved by the introduction of the
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    72
following constant.
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    73
*}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    74
37153
8feed34275ce renamed constant "apply" to "fun_app" (which is closer to the related "fun_upd")
boehmes
parents: 37151
diff changeset
    75
definition fun_app where "fun_app f x = f x"
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    76
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    77
text {*
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    78
Some solvers support a theory of arrays which can be used to encode
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    79
higher-order functions.  The following set of lemmas specifies the
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    80
properties of such (extensional) arrays.
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    81
*}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    82
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    83
lemmas array_rules = ext fun_upd_apply fun_upd_same fun_upd_other
37157
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37153
diff changeset
    84
  fun_upd_upd fun_app_def
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    85
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    86
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    87
36902
c6bae4456741 use 'subsection' instead of 'section', to maintain 1 chapter per file in generated document
huffman
parents: 36899
diff changeset
    88
subsection {* First-order logic *}
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    89
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    90
text {*
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    91
Some SMT solvers require a strict separation between formulas and
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    92
terms.  When translating higher-order into first-order problems,
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    93
all uninterpreted constants (those not builtin in the target solver)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    94
are treated as function symbols in the first-order sense.  Their
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    95
occurrences as head symbols in atoms (i.e., as predicate symbols) is
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    96
turned into terms by equating such atoms with @{term True} using the
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    97
following term-level equation symbol.
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    98
*}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    99
37124
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36902
diff changeset
   100
definition term_eq :: "bool \<Rightarrow> bool \<Rightarrow> bool" where "term_eq x y = (x = y)"
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   101
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   102
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   103
37151
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   104
subsection {* Integer division and modulo for Z3 *}
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   105
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   106
definition z3div :: "int \<Rightarrow> int \<Rightarrow> int" where
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   107
  "z3div k l = (if 0 \<le> l then k div l else -(k div (-l)))"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   108
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   109
definition z3mod :: "int \<Rightarrow> int \<Rightarrow> int" where
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   110
  "z3mod k l = (if 0 \<le> l then k mod l else k mod (-l))"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   111
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   112
lemma div_by_z3div: "k div l = (
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   113
     if k = 0 \<or> l = 0 then 0
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   114
     else if (0 < k \<and> 0 < l) \<or> (k < 0 \<and> 0 < l) then z3div k l
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   115
     else z3div (-k) (-l))"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   116
  by (auto simp add: z3div_def)
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   117
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   118
lemma mod_by_z3mod: "k mod l = (
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   119
     if l = 0 then k
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   120
     else if k = 0 then 0
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   121
     else if (0 < k \<and> 0 < l) \<or> (k < 0 \<and> 0 < l) then z3mod k l
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   122
     else - z3mod (-k) (-l))"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   123
  by (auto simp add: z3mod_def)
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   124
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   125
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   126
36902
c6bae4456741 use 'subsection' instead of 'section', to maintain 1 chapter per file in generated document
huffman
parents: 36899
diff changeset
   127
subsection {* Setup *}
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   128
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   129
use "Tools/SMT/smt_monomorph.ML"
40277
4e3a3461c1a6 added crafted list of SMT built-in constants
boehmes
parents: 40274
diff changeset
   130
use "Tools/SMT/smt_builtin.ML"
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   131
use "Tools/SMT/smt_normalize.ML"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   132
use "Tools/SMT/smt_translate.ML"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   133
use "Tools/SMT/smt_solver.ML"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   134
use "Tools/SMT/smtlib_interface.ML"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   135
use "Tools/SMT/z3_interface.ML"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   136
use "Tools/SMT/z3_proof_parser.ML"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   137
use "Tools/SMT/z3_proof_tools.ML"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   138
use "Tools/SMT/z3_proof_literals.ML"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   139
use "Tools/SMT/z3_proof_reconstruction.ML"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   140
use "Tools/SMT/z3_model.ML"
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   141
use "Tools/SMT/smt_setup_solvers.ML"
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   142
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   143
setup {*
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   144
  SMT_Solver.setup #>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   145
  Z3_Proof_Reconstruction.setup #>
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   146
  SMT_Setup_Solvers.setup
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   147
*}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   148
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   149
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   150
36902
c6bae4456741 use 'subsection' instead of 'section', to maintain 1 chapter per file in generated document
huffman
parents: 36899
diff changeset
   151
subsection {* Configuration *}
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   152
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   153
text {*
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   154
The current configuration can be printed by the command
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   155
@{text smt_status}, which shows the values of most options.
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   156
*}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   157
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   158
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   159
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   160
subsection {* General configuration options *}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   161
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   162
text {*
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   163
The option @{text smt_solver} can be used to change the target SMT
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   164
solver.  The possible values are @{text cvc3}, @{text yices}, and
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   165
@{text z3}.  It is advisable to locally install the selected solver,
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   166
although this is not necessary for @{text cvc3} and @{text z3}, which
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   167
can also be used over an Internet-based service.
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   168
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   169
When using local SMT solvers, the path to their binaries should be
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   170
declared by setting the following environment variables:
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   171
@{text CVC3_SOLVER}, @{text YICES_SOLVER}, and @{text Z3_SOLVER}.
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   172
*}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   173
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   174
declare [[ smt_solver = z3 ]]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   175
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   176
text {*
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   177
Since SMT solvers are potentially non-terminating, there is a timeout
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   178
(given in seconds) to restrict their runtime.  A value greater than
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   179
120 (seconds) is in most cases not advisable.
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   180
*}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   181
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   182
declare [[ smt_timeout = 20 ]]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   183
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   184
text {*
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   185
In general, the binding to SMT solvers runs as an oracle, i.e, the SMT
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   186
solvers are fully trusted without additional checks.  The following
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   187
option can cause the SMT solver to run in proof-producing mode, giving
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   188
a checkable certificate.  This is currently only implemented for Z3.
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   189
*}
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   190
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   191
declare [[ smt_oracle = false ]]
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   192
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   193
text {*
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   194
Each SMT solver provides several commandline options to tweak its
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   195
behaviour.  They can be passed to the solver by setting the following
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   196
options.
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   197
*}
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   198
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   199
declare [[ cvc3_options = "", yices_options = "", z3_options = "" ]]
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   200
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   201
text {*
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   202
Enable the following option to use built-in support for datatypes and
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   203
records.  Currently, this is only implemented for Z3 running in oracle
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   204
mode.
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   205
*}
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   206
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   207
declare [[ smt_datatypes = false ]]
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   208
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   209
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   210
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   211
subsection {* Certificates *}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   212
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   213
text {*
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   214
By setting the option @{text smt_certificates} to the name of a file,
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   215
all following applications of an SMT solver a cached in that file.
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   216
Any further application of the same SMT solver (using the very same
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   217
configuration) re-uses the cached certificate instead of invoking the
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   218
solver.  An empty string disables caching certificates.
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   219
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   220
The filename should be given as an explicit path.  It is good
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   221
practice to use the name of the current theory (with ending
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   222
@{text ".certs"} instead of @{text ".thy"}) as the certificates file.
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   223
*}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   224
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   225
declare [[ smt_certificates = "" ]]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   226
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   227
text {*
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   228
The option @{text smt_fixed} controls whether only stored
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   229
certificates are should be used or invocation of an SMT solver is
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   230
allowed.  When set to @{text true}, no SMT solver will ever be
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   231
invoked and only the existing certificates found in the configured
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   232
cache are used;  when set to @{text false} and there is no cached
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   233
certificate for some proposition, then the configured SMT solver is
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   234
invoked.
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   235
*}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   236
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   237
declare [[ smt_fixed = false ]]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   238
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   239
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   240
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   241
subsection {* Tracing *}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   242
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   243
text {*
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   244
For tracing the generated problem file given to the SMT solver as
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   245
well as the returned result of the solver, the option
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   246
@{text smt_trace} should be set to @{text true}.
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   247
*}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   248
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   249
declare [[ smt_trace = false ]]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   250
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   251
text {*
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   252
From the set of assumptions given to the SMT solver, those assumptions
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   253
used in the proof are traced when the following option is set to
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   254
@{term true}.  This only works for Z3 when it runs in non-oracle mode
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   255
(see options @{text smt_solver} and @{text smt_oracle} above).
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   256
*}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   257
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 39483
diff changeset
   258
declare [[ smt_trace_used_facts = false ]]
39298
5aefb5bc8a93 added preliminary support for SMT datatypes (for now restricted to tuples and lists); only the Z3 interface (in oracle mode) makes use of it, there is especially no Z3 proof reconstruction support for datatypes yet
boehmes
parents: 37818
diff changeset
   259
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   260
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   261
36902
c6bae4456741 use 'subsection' instead of 'section', to maintain 1 chapter per file in generated document
huffman
parents: 36899
diff changeset
   262
subsection {* Schematic rules for Z3 proof reconstruction *}
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   263
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   264
text {*
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   265
Several prof rules of Z3 are not very well documented.  There are two
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   266
lemma groups which can turn failing Z3 proof reconstruction attempts
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   267
into succeeding ones: the facts in @{text z3_rule} are tried prior to
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   268
any implemented reconstruction procedure for all uncertain Z3 proof
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   269
rules;  the facts in @{text z3_simp} are only fed to invocations of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   270
the simplifier when reconstructing theory-specific proof steps.
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   271
*}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   272
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   273
lemmas [z3_rule] =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   274
  refl eq_commute conj_commute disj_commute simp_thms nnf_simps
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   275
  ring_distribs field_simps times_divide_eq_right times_divide_eq_left
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   276
  if_True if_False not_not
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   277
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   278
lemma [z3_rule]:
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   279
  "(P \<longrightarrow> Q) = (Q \<or> \<not>P)"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   280
  "(\<not>P \<longrightarrow> Q) = (P \<or> Q)"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   281
  "(\<not>P \<longrightarrow> Q) = (Q \<or> P)"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   282
  by auto
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   283
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   284
lemma [z3_rule]:
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   285
  "((P = Q) \<longrightarrow> R) = (R | (Q = (\<not>P)))"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   286
  by auto
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   287
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   288
lemma [z3_rule]:
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   289
  "((\<not>P) = P) = False"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   290
  "(P = (\<not>P)) = False"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   291
  "(P \<noteq> Q) = (Q = (\<not>P))"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   292
  "(P = Q) = ((\<not>P \<or> Q) \<and> (P \<or> \<not>Q))"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   293
  "(P \<noteq> Q) = ((\<not>P \<or> \<not>Q) \<and> (P \<or> Q))"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   294
  by auto
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   295
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   296
lemma [z3_rule]:
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   297
  "(if P then P else \<not>P) = True"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   298
  "(if \<not>P then \<not>P else P) = True"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   299
  "(if P then True else False) = P"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   300
  "(if P then False else True) = (\<not>P)"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   301
  "(if \<not>P then x else y) = (if P then y else x)"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   302
  by auto
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   303
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   304
lemma [z3_rule]:
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   305
  "P = Q \<or> P \<or> Q"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   306
  "P = Q \<or> \<not>P \<or> \<not>Q"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   307
  "(\<not>P) = Q \<or> \<not>P \<or> Q"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   308
  "(\<not>P) = Q \<or> P \<or> \<not>Q"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   309
  "P = (\<not>Q) \<or> \<not>P \<or> Q"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   310
  "P = (\<not>Q) \<or> P \<or> \<not>Q"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   311
  "P \<noteq> Q \<or> P \<or> \<not>Q"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   312
  "P \<noteq> Q \<or> \<not>P \<or> Q"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   313
  "P \<noteq> (\<not>Q) \<or> P \<or> Q"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   314
  "(\<not>P) \<noteq> Q \<or> P \<or> Q"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   315
  "P \<or> Q \<or> P \<noteq> (\<not>Q)"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   316
  "P \<or> Q \<or> (\<not>P) \<noteq> Q"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   317
  "P \<or> \<not>Q \<or> P \<noteq> Q"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   318
  "\<not>P \<or> Q \<or> P \<noteq> Q"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   319
  by auto
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   320
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   321
lemma [z3_rule]:
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   322
  "0 + (x::int) = x"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   323
  "x + 0 = x"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   324
  "0 * x = 0"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   325
  "1 * x = x"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   326
  "x + y = y + x"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   327
  by auto
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   328
37124
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36902
diff changeset
   329
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36902
diff changeset
   330
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36902
diff changeset
   331
hide_type (open) pattern
37153
8feed34275ce renamed constant "apply" to "fun_app" (which is closer to the related "fun_upd")
boehmes
parents: 37151
diff changeset
   332
hide_const Pattern term_eq
40274
6486c610a549 introduced SMT.distinct as a representation of the solvers' built-in predicate; check that SMT.distinct is always applied to an explicit list
boehmes
parents: 40162
diff changeset
   333
hide_const (open) trigger pat nopat distinct fun_app z3div z3mod
37124
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36902
diff changeset
   334
39483
9f0e5684f04b add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
boehmes
parents: 39298
diff changeset
   335
9f0e5684f04b add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
boehmes
parents: 39298
diff changeset
   336
9f0e5684f04b add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
boehmes
parents: 39298
diff changeset
   337
subsection {* Selectors for datatypes *}
9f0e5684f04b add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
boehmes
parents: 39298
diff changeset
   338
9f0e5684f04b add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
boehmes
parents: 39298
diff changeset
   339
setup {* Datatype_Selectors.setup *}
9f0e5684f04b add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
boehmes
parents: 39298
diff changeset
   340
9f0e5684f04b add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
boehmes
parents: 39298
diff changeset
   341
declare [[ selector Pair 1 = fst, selector Pair 2 = snd ]]
9f0e5684f04b add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
boehmes
parents: 39298
diff changeset
   342
declare [[ selector Cons 1 = hd, selector Cons 2 = tl ]]
9f0e5684f04b add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
boehmes
parents: 39298
diff changeset
   343
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   344
end