src/HOL/SMT_Examples/SMT_Tests.thy
author blanchet
Thu, 01 Nov 2012 11:34:00 +0100
changeset 49995 3b7ad6153322
parent 49834 b27bbb021df1
child 50662 b1f4291eb916
permissions -rw-r--r--
regenerated "SMT_Examples" certificates after soft-timeout change + removed a few needless oracles
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_Tests.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 {* Tests for the 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_Tests
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
     8
imports Complex_Main
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
47152
446cfc760ccf renamed "smt_fixed" to "smt_read_only_certificates"
blanchet
parents: 47111
diff changeset
    11
declare [[smt_oracle = false]]
446cfc760ccf renamed "smt_fixed" to "smt_read_only_certificates"
blanchet
parents: 47111
diff changeset
    12
declare [[smt_certificates = "SMT_Tests.certs"]]
446cfc760ccf renamed "smt_fixed" to "smt_read_only_certificates"
blanchet
parents: 47111
diff changeset
    13
declare [[smt_read_only_certificates = true]]
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    14
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
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    17
smt_status
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    18
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
text {* Most examples are taken from various Isabelle theories and from HOL4. *}
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    22
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    23
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    24
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    25
section {* Propositional logic *}
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    26
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    27
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    28
  "True"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    29
  "\<not>False"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    30
  "\<not>\<not>True"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    31
  "True \<and> True"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    32
  "True \<or> False"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    33
  "False \<longrightarrow> True"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    34
  "\<not>(False \<longleftrightarrow> True)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    35
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    36
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    37
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    38
  "P \<or> \<not>P"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    39
  "\<not>(P \<and> \<not>P)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    40
  "(True \<and> P) \<or> \<not>P \<or> (False \<and> P) \<or> P"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    41
  "P \<longrightarrow> P"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    42
  "P \<and> \<not> P \<longrightarrow> False"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    43
  "P \<and> Q \<longrightarrow> Q \<and> P"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    44
  "P \<or> Q \<longrightarrow> Q \<or> P"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    45
  "P \<and> Q \<longrightarrow> P \<or> Q"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    46
  "\<not>(P \<or> Q) \<longrightarrow> \<not>P"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    47
  "\<not>(P \<or> Q) \<longrightarrow> \<not>Q"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    48
  "\<not>P \<longrightarrow> \<not>(P \<and> Q)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    49
  "\<not>Q \<longrightarrow> \<not>(P \<and> Q)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    50
  "(P \<and> Q) \<longleftrightarrow> (\<not>(\<not>P \<or> \<not>Q))"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    51
  "(P \<and> Q) \<and> R \<longrightarrow> P \<and> (Q \<and> R)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    52
  "(P \<or> Q) \<or> R \<longrightarrow> P \<or> (Q \<or> R)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    53
  "(P \<and> Q) \<or> R  \<longrightarrow> (P \<or> R) \<and> (Q \<or> R)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    54
  "(P \<or> R) \<and> (Q \<or> R) \<longrightarrow> (P \<and> Q) \<or> R"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    55
  "(P \<or> Q) \<and> R \<longrightarrow> (P \<and> R) \<or> (Q \<and> R)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    56
  "(P \<and> R) \<or> (Q \<and> R) \<longrightarrow> (P \<or> Q) \<and> R"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    57
  "((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    58
  "(P \<longrightarrow> R) \<and> (Q \<longrightarrow> R) \<longleftrightarrow> (P \<or> Q \<longrightarrow> R)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    59
  "(P \<and> Q \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> (Q \<longrightarrow> R))"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    60
  "((P \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow>  ((Q \<longrightarrow> R) \<longrightarrow> R) \<longrightarrow> (P \<and> Q \<longrightarrow> R) \<longrightarrow> R"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    61
  "\<not>(P \<longrightarrow> R) \<longrightarrow>  \<not>(Q \<longrightarrow> R) \<longrightarrow> \<not>(P \<and> Q \<longrightarrow> R)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    62
  "(P \<longrightarrow> Q \<and> R) \<longleftrightarrow> (P \<longrightarrow> Q) \<and> (P \<longrightarrow> R)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    63
  "P \<longrightarrow> (Q \<longrightarrow> P)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    64
  "(P \<longrightarrow> Q \<longrightarrow> R) \<longrightarrow> (P \<longrightarrow> Q)\<longrightarrow> (P \<longrightarrow> R)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    65
  "(P \<longrightarrow> Q) \<or> (P \<longrightarrow> R) \<longrightarrow> (P \<longrightarrow> Q \<or> R)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    66
  "((((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P) \<longrightarrow> Q) \<longrightarrow> Q"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    67
  "(P \<longrightarrow> Q) \<longrightarrow> (\<not>Q \<longrightarrow> \<not>P)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    68
  "(P \<longrightarrow> Q \<or> R) \<longrightarrow> (P \<longrightarrow> Q) \<or> (P \<longrightarrow> R)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    69
  "(P \<longrightarrow> Q) \<and> (Q  \<longrightarrow> P) \<longrightarrow> (P \<longleftrightarrow> Q)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    70
  "(P \<longleftrightarrow> Q) \<longleftrightarrow> (Q \<longleftrightarrow> P)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    71
  "\<not>(P \<longleftrightarrow> \<not>P)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    72
  "(P \<longrightarrow> Q) \<longleftrightarrow> (\<not>Q \<longrightarrow> \<not>P)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    73
  "P \<longleftrightarrow> P \<longleftrightarrow> P \<longleftrightarrow> P \<longleftrightarrow> P \<longleftrightarrow> P \<longleftrightarrow> P \<longleftrightarrow> P \<longleftrightarrow> P \<longleftrightarrow> P"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    74
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    75
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    76
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    77
  "(if P then Q1 else Q2) \<longleftrightarrow> ((P \<longrightarrow> Q1) \<and> (\<not>P \<longrightarrow> Q2))"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    78
  "if P then (Q \<longrightarrow> P) else (P \<longrightarrow> Q)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    79
  "(if P1 \<or> P2 then Q1 else Q2) \<longleftrightarrow> (if P1 then Q1 else if P2 then Q1 else Q2)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    80
  "(if P1 \<and> P2 then Q1 else Q2) \<longleftrightarrow> (if P1 then if P2 then Q1 else Q2 else Q2)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    81
  "(P1 \<longrightarrow> (if P2 then Q1 else Q2)) \<longleftrightarrow>
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    82
   (if P1 \<longrightarrow> P2 then P1 \<longrightarrow> Q1 else P1 \<longrightarrow> Q2)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    83
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    84
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    85
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    86
  "case P of True \<Rightarrow> P | False \<Rightarrow> \<not>P"
37786
4eb98849c5c0 fixed handling of Ball/Bex: turn equalities into meta-equalities for the rewriting conversions;
boehmes
parents: 37157
diff changeset
    87
  "case P of False \<Rightarrow> \<not>P | True \<Rightarrow> P"
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    88
  "case \<not>P of True \<Rightarrow> \<not>P | False \<Rightarrow> P"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    89
  "case P of True \<Rightarrow> (Q \<longrightarrow> P) | False \<Rightarrow> (P \<longrightarrow> Q)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    90
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    91
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    92
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    93
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    94
section {* First-order logic with equality *}
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    95
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    96
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    97
  "x = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    98
  "x = y \<longrightarrow> y = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
    99
  "x = y \<and> y = z \<longrightarrow> x = z"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   100
  "x = y \<longrightarrow> f x = f y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   101
  "x = y \<longrightarrow> g x y = g y x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   102
  "f (f x) = x \<and> f (f (f (f (f x)))) = x \<longrightarrow> f x = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   103
  "((if a then b else c) = d) = ((a \<longrightarrow> (b = d)) \<and> (\<not> a \<longrightarrow> (c = d)))"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   104
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   105
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   106
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   107
  "\<forall>x. x = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   108
  "(\<forall>x. P x) \<longleftrightarrow> (\<forall>y. P y)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   109
  "\<forall>x. P x \<longrightarrow> (\<forall>y. P x \<or> P y)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   110
  "(\<forall>x. P x \<and> Q x) \<longleftrightarrow> (\<forall>x. P x) \<and> (\<forall>x. Q x)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   111
  "(\<forall>x. P x) \<or> R \<longleftrightarrow> (\<forall>x. P x \<or> R)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   112
  "(\<forall>x. P x) \<and> R \<longleftrightarrow> (\<forall>x. P x \<and> R)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   113
  "(\<forall>x y z. S x z) \<longleftrightarrow> (\<forall>x z. S x z)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   114
  "(\<forall>x y. S x y \<longrightarrow> S y x) \<longrightarrow> (\<forall>x. S x y) \<longrightarrow> S y x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   115
  "(\<forall>x. P x \<longrightarrow> P (f x)) \<and> P d \<longrightarrow> P (f(f(f(d))))"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   116
  "(\<forall>x y. s x y = s y x) \<longrightarrow> a = a \<and> s a b = s b a"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   117
  "(\<forall>s. q s \<longrightarrow> r s) \<and> \<not>r s \<and> (\<forall>s. \<not>r s \<and> \<not>q s \<longrightarrow> p t \<or> q t) \<longrightarrow> p t \<or> r t"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   118
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   119
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   120
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   121
  "\<exists>x. x = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   122
  "(\<exists>x. P x) \<longleftrightarrow> (\<exists>y. P y)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   123
  "(\<exists>x. P x \<or> Q x) \<longleftrightarrow> (\<exists>x. P x) \<or> (\<exists>x. Q x)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   124
  "(\<exists>x. P x) \<and> R \<longleftrightarrow> (\<exists>x. P x \<and> R)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   125
  "(\<exists>x y z. S x z) \<longleftrightarrow> (\<exists>x z. S x z)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   126
  "\<not>((\<exists>x. \<not>P x) \<and> ((\<exists>x. P x) \<or> (\<exists>x. P x \<and> Q x)) \<and> \<not>(\<exists>x. P x))"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   127
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   128
49995
3b7ad6153322 regenerated "SMT_Examples" certificates after soft-timeout change + removed a few needless oracles
blanchet
parents: 49834
diff changeset
   129
lemma
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   130
  "\<exists>x y. x = y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   131
  "\<exists>x. P x \<longrightarrow> (\<exists>y. P x \<and> P y)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   132
  "(\<exists>x. P x) \<or> R \<longleftrightarrow> (\<exists>x. P x \<or> R)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   133
  "\<exists>x. P x \<longrightarrow> P a \<and> P b"
49995
3b7ad6153322 regenerated "SMT_Examples" certificates after soft-timeout change + removed a few needless oracles
blanchet
parents: 49834
diff changeset
   134
  "\<exists>x. (\<exists>y. P y) \<longrightarrow> P x"
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   135
  "(\<exists>x. Q \<longrightarrow> P x) \<longleftrightarrow> (Q \<longrightarrow> (\<exists>x. P x))"
49995
3b7ad6153322 regenerated "SMT_Examples" certificates after soft-timeout change + removed a few needless oracles
blanchet
parents: 49834
diff changeset
   136
by smt+
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   137
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   138
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   139
  "(\<not>(\<exists>x. P x)) \<longleftrightarrow> (\<forall>x. \<not> P x)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   140
  "(\<exists>x. P x \<longrightarrow> Q) \<longleftrightarrow> (\<forall>x. P x) \<longrightarrow> Q"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   141
  "(\<forall>x y. R x y = x) \<longrightarrow> (\<exists>y. R x y) = R x c"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   142
  "(if P x then \<not>(\<exists>y. P y) else (\<forall>y. \<not>P y)) \<longrightarrow> P x \<longrightarrow> P y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   143
  "(\<forall>x y. R x y = x) \<and> (\<forall>x. \<exists>y. R x y) = (\<forall>x. R x c) \<longrightarrow> (\<exists>y. R x y) = R x c"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   144
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   145
49995
3b7ad6153322 regenerated "SMT_Examples" certificates after soft-timeout change + removed a few needless oracles
blanchet
parents: 49834
diff changeset
   146
lemma
41132
42384824b732 updated SMT certificates
boehmes
parents: 40681
diff changeset
   147
  "\<forall>x. \<exists>y. f x y = f x (g x)"
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   148
  "(\<not>\<not>(\<exists>x. P x)) \<longleftrightarrow> (\<not>(\<forall>x. \<not> P x))"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   149
  "\<forall>u. \<exists>v. \<forall>w. \<exists>x. f u v w x = f u (g u) w (h u w)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   150
  "\<exists>x. if x = y then (\<forall>y. y = x \<or> y \<noteq> x) else (\<forall>y. y = (x, x) \<or> y \<noteq> (x, x))"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   151
  "\<exists>x. if x = y then (\<exists>y. y = x \<or> y \<noteq> x) else (\<exists>y. y = (x, x) \<or> y \<noteq> (x, x))"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   152
  "(\<exists>x. \<forall>y. P x \<longleftrightarrow> P y) \<longrightarrow> ((\<exists>x. P x) \<longleftrightarrow> (\<forall>y. P y))"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   153
  "\<exists>z. P z \<longrightarrow> (\<forall>x. P x)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   154
  "(\<exists>y. \<forall>x. R x y) \<longrightarrow> (\<forall>x. \<exists>y. R x y)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   155
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   156
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   157
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   158
  "(\<exists>! x. P x) \<longrightarrow> (\<exists>x. P x)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   159
  "(\<exists>!x. P x) \<longleftrightarrow> (\<exists>x. P x \<and> (\<forall>y. y \<noteq> x \<longrightarrow> \<not>P y))"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   160
  "P a \<longrightarrow> (\<forall>x. P x \<longrightarrow> x = a) \<longrightarrow> (\<exists>!x. P x)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   161
  "(\<exists>x. P x) \<and> (\<forall>x y. P x \<and> P y \<longrightarrow> x = y) \<longrightarrow> (\<exists>!x. P x)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   162
  "(\<exists>!x. P x) \<and> (\<forall>x. P x \<and> (\<forall>y. P y \<longrightarrow> y = x) \<longrightarrow> R) \<longrightarrow> R"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   163
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   164
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   165
lemma
37786
4eb98849c5c0 fixed handling of Ball/Bex: turn equalities into meta-equalities for the rewriting conversions;
boehmes
parents: 37157
diff changeset
   166
  "(\<forall>x\<in>M. P x) \<and> c \<in> M \<longrightarrow> P c"
4eb98849c5c0 fixed handling of Ball/Bex: turn equalities into meta-equalities for the rewriting conversions;
boehmes
parents: 37157
diff changeset
   167
  "(\<exists>x\<in>M. P x) \<or> \<not>(P c \<and> c \<in> M)"
4eb98849c5c0 fixed handling of Ball/Bex: turn equalities into meta-equalities for the rewriting conversions;
boehmes
parents: 37157
diff changeset
   168
  by smt+
4eb98849c5c0 fixed handling of Ball/Bex: turn equalities into meta-equalities for the rewriting conversions;
boehmes
parents: 37157
diff changeset
   169
4eb98849c5c0 fixed handling of Ball/Bex: turn equalities into meta-equalities for the rewriting conversions;
boehmes
parents: 37157
diff changeset
   170
lemma
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   171
  "let P = True in P"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   172
  "let P = P1 \<or> P2 in P \<or> \<not>P"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   173
  "let P1 = True; P2 = False in P1 \<and> P2 \<longrightarrow> P2 \<or> P1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   174
  "(let x = y in x) = y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   175
  "(let x = y in Q x) \<longleftrightarrow> (let z = y in Q z)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   176
  "(let x = y1; z = y2 in R x z) \<longleftrightarrow> (let z = y2; x = y1 in R x z)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   177
  "(let x = y1; z = y2 in R x z) \<longleftrightarrow> (let z = y1; x = y2 in R z x)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   178
  "let P = (\<forall>x. Q x) in if P then P else \<not>P"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   179
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   180
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   181
lemma
47155
ade3fc826af3 dropped support for List.distinct in binding to SMT solvers: only few applications benefited from this support, and in some cases the smt method fails due to its support for List.distinct
boehmes
parents: 47152
diff changeset
   182
  "a \<noteq> b \<and> a \<noteq> c \<and> b \<noteq> c \<and> (\<forall>x y. f x = f y \<longrightarrow> y = x) \<longrightarrow> f a \<noteq> f b"
41899
83dd157ec9ab finished and tested Z3 proof reconstruction for injective functions;
boehmes
parents: 41601
diff changeset
   183
  by smt
83dd157ec9ab finished and tested Z3 proof reconstruction for injective functions;
boehmes
parents: 41601
diff changeset
   184
83dd157ec9ab finished and tested Z3 proof reconstruction for injective functions;
boehmes
parents: 41601
diff changeset
   185
lemma
83dd157ec9ab finished and tested Z3 proof reconstruction for injective functions;
boehmes
parents: 41601
diff changeset
   186
  "(\<forall>x y z. f x y = f x z \<longrightarrow> y = z) \<and> b \<noteq> c \<longrightarrow> f a b \<noteq> f a c"
83dd157ec9ab finished and tested Z3 proof reconstruction for injective functions;
boehmes
parents: 41601
diff changeset
   187
  "(\<forall>x y z. f x y = f z y \<longrightarrow> x = z) \<and> a \<noteq> d \<longrightarrow> f a b \<noteq> f d b"
83dd157ec9ab finished and tested Z3 proof reconstruction for injective functions;
boehmes
parents: 41601
diff changeset
   188
  by smt+
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   189
44753
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   190
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   191
section {* Guidance for quantifier heuristics: patterns and weights *}
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   192
37124
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   193
lemma
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   194
  assumes "\<forall>x. SMT.trigger [[SMT.pat (f x)]] (f x = x)"
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   195
  shows "f 1 = 1"
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   196
  using assms by smt
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   197
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   198
lemma
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   199
  assumes "\<forall>x y. SMT.trigger [[SMT.pat (f x), SMT.pat (g y)]] (f x = g y)"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45694
diff changeset
   200
  shows "f a = g b"
47111
a4476e55a241 reintroduced broken proofs and regenerated certificates
blanchet
parents: 47108
diff changeset
   201
  using assms by smt
37124
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   202
44753
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   203
lemma
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   204
  assumes "ALL x. SMT.trigger [[SMT.pat (P x)]] (P x --> Q x)"
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   205
  and "P t"
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   206
  shows "Q t"
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   207
  using assms by smt
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   208
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   209
lemma
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   210
  assumes "ALL x. SMT.trigger [[SMT.pat (P x), SMT.pat (Q x)]]
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   211
    (P x & Q x --> R x)"
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   212
  and "P t" and "Q t"
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   213
  shows "R t"
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   214
  using assms by smt
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   215
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   216
lemma
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   217
  assumes "ALL x. SMT.trigger [[SMT.pat (P x)], [SMT.pat (Q x)]]
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   218
    ((P x --> R x) & (Q x --> R x))"
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   219
  and "P t | Q t"
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   220
  shows "R t"
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   221
  using assms by smt
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   222
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   223
lemma
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   224
  assumes "ALL x. SMT.trigger [[SMT.pat (P x)]] (SMT.weight 2 (P x --> Q x))"
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   225
  and "P t"
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   226
  shows "Q t"
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   227
  using assms by smt
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   228
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   229
lemma
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   230
  assumes "ALL x. SMT.weight 1 (P x --> Q x)"
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   231
  and "P t"
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   232
  shows "Q t"
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   233
  using assms by smt
3c73f4068978 added some examples for pattern and weight annotations
boehmes
parents: 41899
diff changeset
   234
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   235
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   236
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   237
section {* Meta logical connectives *}
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   238
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   239
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   240
  "True \<Longrightarrow> True"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   241
  "False \<Longrightarrow> True"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   242
  "False \<Longrightarrow> False"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   243
  "P' x \<Longrightarrow> P' x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   244
  "P \<Longrightarrow> P \<or> Q"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   245
  "Q \<Longrightarrow> P \<or> Q"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   246
  "\<not>P \<Longrightarrow> P \<longrightarrow> Q"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   247
  "Q \<Longrightarrow> P \<longrightarrow> Q"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   248
  "\<lbrakk>P; \<not>Q\<rbrakk> \<Longrightarrow> \<not>(P \<longrightarrow> Q)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   249
  "P' x \<equiv> P' x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   250
  "P' x \<equiv> Q' x \<Longrightarrow> P' x = Q' x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   251
  "P' x = Q' x \<Longrightarrow> P' x \<equiv> Q' x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   252
  "x \<equiv> y \<Longrightarrow> y \<equiv> z \<Longrightarrow> x \<equiv> (z::'a::type)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   253
  "x \<equiv> y \<Longrightarrow> (f x :: 'b::type) \<equiv> f y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   254
  "(\<And>x. g x) \<Longrightarrow> g a \<or> a"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   255
  "(\<And>x y. h x y \<and> h y x) \<Longrightarrow> \<forall>x. h x x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   256
  "(p \<or> q) \<and> \<not>p \<Longrightarrow> q"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   257
  "(a \<and> b) \<or> (c \<and> d) \<Longrightarrow> (a \<and> b) \<or> (c \<and> d)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   258
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   259
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   260
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   261
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   262
section {* Natural numbers *}
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   263
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   264
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   265
  "(0::nat) = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   266
  "(1::nat) = 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   267
  "(0::nat) < 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   268
  "(0::nat) \<le> 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   269
  "(123456789::nat) < 2345678901"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   270
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   271
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   272
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   273
  "Suc 0 = 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   274
  "Suc x = x + 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   275
  "x < Suc x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   276
  "(Suc x = Suc y) = (x = y)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   277
  "Suc (x + y) < Suc x + Suc y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   278
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   279
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   280
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   281
  "(x::nat) + 0 = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   282
  "0 + x = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   283
  "x + y = y + x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   284
  "x + (y + z) = (x + y) + z"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   285
  "(x + y = 0) = (x = 0 \<and> y = 0)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   286
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   287
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   288
lemma 
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   289
  "(x::nat) - 0 = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   290
  "x < y \<longrightarrow> x - y = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   291
  "x - y = 0 \<or> y - x = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   292
  "(x - y) + y = (if x < y then y else x)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   293
  "x - y - z = x - (y + z)" 
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   294
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   295
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   296
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   297
  "(x::nat) * 0 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   298
  "0 * x = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   299
  "x * 1 = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   300
  "1 * x = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   301
  "3 * x = x * 3"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   302
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   303
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   304
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   305
  "(0::nat) div 0 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   306
  "(x::nat) div 0 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   307
  "(0::nat) div 1 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   308
  "(1::nat) div 1 = 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   309
  "(3::nat) div 1 = 3"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   310
  "(x::nat) div 1 = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   311
  "(0::nat) div 3 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   312
  "(1::nat) div 3 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   313
  "(3::nat) div 3 = 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   314
  "(x::nat) div 3 \<le> x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   315
  "(x div 3 = x) = (x = 0)"
48069
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 47155
diff changeset
   316
  using [[z3_with_extensions]]
37151
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   317
  by smt+
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   318
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   319
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   320
  "(0::nat) mod 0 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   321
  "(x::nat) mod 0 = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   322
  "(0::nat) mod 1 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   323
  "(1::nat) mod 1 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   324
  "(3::nat) mod 1 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   325
  "(x::nat) mod 1 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   326
  "(0::nat) mod 3 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   327
  "(1::nat) mod 3 = 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   328
  "(3::nat) mod 3 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   329
  "x mod 3 < 3"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   330
  "(x mod 3 = x) = (x < 3)"
48069
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 47155
diff changeset
   331
  using [[z3_with_extensions]]
37151
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   332
  by smt+
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   333
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   334
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   335
  "(x::nat) = x div 1 * 1 + x mod 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   336
  "x = x div 3 * 3 + x mod 3"
48069
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 47155
diff changeset
   337
  using [[z3_with_extensions]]
37151
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   338
  by smt+
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   339
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   340
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   341
  "min (x::nat) y \<le> x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   342
  "min x y \<le> y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   343
  "min x y \<le> x + y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   344
  "z < x \<and> z < y \<longrightarrow> z < min x y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   345
  "min x y = min y x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   346
  "min x 0 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   347
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   348
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   349
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   350
  "max (x::nat) y \<ge> x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   351
  "max x y \<ge> y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   352
  "max x y \<ge> (x - y) + (y - x)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   353
  "z > x \<and> z > y \<longrightarrow> z > max x y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   354
  "max x y = max y x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   355
  "max x 0 = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   356
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   357
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   358
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   359
  "0 \<le> (x::nat)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   360
  "0 < x \<and> x \<le> 1 \<longrightarrow> x = 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   361
  "x \<le> x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   362
  "x \<le> y \<longrightarrow> 3 * x \<le> 3 * y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   363
  "x < y \<longrightarrow> 3 * x < 3 * y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   364
  "x < y \<longrightarrow> x \<le> y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   365
  "(x < y) = (x + 1 \<le> y)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   366
  "\<not>(x < x)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   367
  "x \<le> y \<longrightarrow> y \<le> z \<longrightarrow> x \<le> z"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   368
  "x < y \<longrightarrow> y \<le> z \<longrightarrow> x \<le> z"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   369
  "x \<le> y \<longrightarrow> y < z \<longrightarrow> x \<le> z"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   370
  "x < y \<longrightarrow> y < z \<longrightarrow> x < z"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   371
  "x < y \<and> y < z \<longrightarrow> \<not>(z < x)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   372
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   373
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   374
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   375
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   376
section {* Integers *}
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   377
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   378
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   379
  "(0::int) = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   380
  "(0::int) = -0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   381
  "(0::int) = (- 0)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   382
  "(1::int) = 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   383
  "\<not>(-1 = (1::int))"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   384
  "(0::int) < 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   385
  "(0::int) \<le> 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   386
  "-123 + 345 < (567::int)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   387
  "(123456789::int) < 2345678901"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   388
  "(-123456789::int) < 2345678901"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   389
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   390
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   391
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   392
  "(x::int) + 0 = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   393
  "0 + x = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   394
  "x + y = y + x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   395
  "x + (y + z) = (x + y) + z"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   396
  "(x + y = 0) = (x = -y)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   397
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   398
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   399
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   400
  "(-1::int) = - 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   401
  "(-3::int) = - 3"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   402
  "-(x::int) < 0 \<longleftrightarrow> x > 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   403
  "x > 0 \<longrightarrow> -x < 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   404
  "x < 0 \<longrightarrow> -x > 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   405
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   406
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   407
lemma 
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   408
  "(x::int) - 0 = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   409
  "0 - x = -x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   410
  "x < y \<longrightarrow> x - y < 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   411
  "x - y = -(y - x)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   412
  "x - y = -y + x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   413
  "x - y - z = x - (y + z)" 
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   414
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   415
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   416
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   417
  "(x::int) * 0 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   418
  "0 * x = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   419
  "x * 1 = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   420
  "1 * x = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   421
  "x * -1 = -x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   422
  "-1 * x = -x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   423
  "3 * x = x * 3"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   424
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   425
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   426
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   427
  "(0::int) div 0 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   428
  "(x::int) div 0 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   429
  "(0::int) div 1 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   430
  "(1::int) div 1 = 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   431
  "(3::int) div 1 = 3"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   432
  "(x::int) div 1 = x"
37151
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   433
  "(0::int) div -1 = 0"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   434
  "(1::int) div -1 = -1"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   435
  "(3::int) div -1 = -3"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   436
  "(x::int) div -1 = -x"
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   437
  "(0::int) div 3 = 0"
37151
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   438
  "(0::int) div -3 = 0"
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   439
  "(1::int) div 3 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   440
  "(3::int) div 3 = 1"
37151
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   441
  "(5::int) div 3 = 1"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   442
  "(1::int) div -3 = -1"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   443
  "(3::int) div -3 = -1"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   444
  "(5::int) div -3 = -2"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   445
  "(-1::int) div 3 = -1"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   446
  "(-3::int) div 3 = -1"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   447
  "(-5::int) div 3 = -2"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   448
  "(-1::int) div -3 = 0"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   449
  "(-3::int) div -3 = 1"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   450
  "(-5::int) div -3 = 1"
48069
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 47155
diff changeset
   451
  using [[z3_with_extensions]]
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   452
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   453
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   454
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   455
  "(0::int) mod 0 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   456
  "(x::int) mod 0 = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   457
  "(0::int) mod 1 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   458
  "(1::int) mod 1 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   459
  "(3::int) mod 1 = 0"
37151
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   460
  "(x::int) mod 1 = 0"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   461
  "(0::int) mod -1 = 0"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   462
  "(1::int) mod -1 = 0"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   463
  "(3::int) mod -1 = 0"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   464
  "(x::int) mod -1 = 0"
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   465
  "(0::int) mod 3 = 0"
37151
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   466
  "(0::int) mod -3 = 0"
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   467
  "(1::int) mod 3 = 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   468
  "(3::int) mod 3 = 0"
37151
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   469
  "(5::int) mod 3 = 2"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   470
  "(1::int) mod -3 = -2"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   471
  "(3::int) mod -3 = 0"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   472
  "(5::int) mod -3 = -1"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   473
  "(-1::int) mod 3 = 2"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   474
  "(-3::int) mod 3 = 0"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   475
  "(-5::int) mod 3 = 1"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   476
  "(-1::int) mod -3 = -1"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   477
  "(-3::int) mod -3 = 0"
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   478
  "(-5::int) mod -3 = -2"
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   479
  "x mod 3 < 3"
37151
3e9e8dfb3c98 use Z3's builtin support for div and mod
boehmes
parents: 37124
diff changeset
   480
  "(x mod 3 = x) \<longrightarrow> (x < 3)"
48069
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 47155
diff changeset
   481
  using [[z3_with_extensions]]
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   482
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   483
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   484
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   485
  "(x::int) = x div 1 * 1 + x mod 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   486
  "x = x div 3 * 3 + x mod 3"
48069
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 47155
diff changeset
   487
  using [[z3_with_extensions]]
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   488
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   489
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   490
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   491
  "abs (x::int) \<ge> 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   492
  "(abs x = 0) = (x = 0)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   493
  "(x \<ge> 0) = (abs x = x)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   494
  "(x \<le> 0) = (abs x = -x)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   495
  "abs (abs x) = abs x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   496
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   497
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   498
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   499
  "min (x::int) y \<le> x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   500
  "min x y \<le> y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   501
  "z < x \<and> z < y \<longrightarrow> z < min x y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   502
  "min x y = min y x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   503
  "x \<ge> 0 \<longrightarrow> min x 0 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   504
  "min x y \<le> abs (x + y)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   505
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   506
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   507
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   508
  "max (x::int) y \<ge> x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   509
  "max x y \<ge> y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   510
  "z > x \<and> z > y \<longrightarrow> z > max x y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   511
  "max x y = max y x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   512
  "x \<ge> 0 \<longrightarrow> max x 0 = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   513
  "max x y \<ge> - abs x - abs y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   514
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   515
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   516
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   517
  "0 < (x::int) \<and> x \<le> 1 \<longrightarrow> x = 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   518
  "x \<le> x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   519
  "x \<le> y \<longrightarrow> 3 * x \<le> 3 * y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   520
  "x < y \<longrightarrow> 3 * x < 3 * y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   521
  "x < y \<longrightarrow> x \<le> y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   522
  "(x < y) = (x + 1 \<le> y)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   523
  "\<not>(x < x)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   524
  "x \<le> y \<longrightarrow> y \<le> z \<longrightarrow> x \<le> z"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   525
  "x < y \<longrightarrow> y \<le> z \<longrightarrow> x \<le> z"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   526
  "x \<le> y \<longrightarrow> y < z \<longrightarrow> x \<le> z"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   527
  "x < y \<longrightarrow> y < z \<longrightarrow> x < z"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   528
  "x < y \<and> y < z \<longrightarrow> \<not>(z < x)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   529
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   530
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   531
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   532
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   533
section {* Reals *}
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   534
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   535
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   536
  "(0::real) = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   537
  "(0::real) = -0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   538
  "(0::real) = (- 0)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   539
  "(1::real) = 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   540
  "\<not>(-1 = (1::real))"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   541
  "(0::real) < 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   542
  "(0::real) \<le> 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   543
  "-123 + 345 < (567::real)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   544
  "(123456789::real) < 2345678901"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   545
  "(-123456789::real) < 2345678901"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   546
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   547
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   548
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   549
  "(x::real) + 0 = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   550
  "0 + x = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   551
  "x + y = y + x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   552
  "x + (y + z) = (x + y) + z"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   553
  "(x + y = 0) = (x = -y)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   554
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   555
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   556
lemma
41132
42384824b732 updated SMT certificates
boehmes
parents: 40681
diff changeset
   557
  "(-1::real) = - 1"
42384824b732 updated SMT certificates
boehmes
parents: 40681
diff changeset
   558
  "(-3::real) = - 3"
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   559
  "-(x::real) < 0 \<longleftrightarrow> x > 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   560
  "x > 0 \<longrightarrow> -x < 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   561
  "x < 0 \<longrightarrow> -x > 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   562
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   563
49995
3b7ad6153322 regenerated "SMT_Examples" certificates after soft-timeout change + removed a few needless oracles
blanchet
parents: 49834
diff changeset
   564
lemma
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   565
  "(x::real) - 0 = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   566
  "0 - x = -x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   567
  "x < y \<longrightarrow> x - y < 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   568
  "x - y = -(y - x)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   569
  "x - y = -y + x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   570
  "x - y - z = x - (y + z)" 
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   571
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   572
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   573
lemma
41132
42384824b732 updated SMT certificates
boehmes
parents: 40681
diff changeset
   574
  "(x::real) * 0 = 0"
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   575
  "0 * x = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   576
  "x * 1 = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   577
  "1 * x = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   578
  "x * -1 = -x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   579
  "-1 * x = -x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   580
  "3 * x = x * 3"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   581
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   582
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   583
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   584
  "(1/2 :: real) < 1"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   585
  "(1::real) / 3 = 1 / 3"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   586
  "(1::real) / -3 = - 1 / 3"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   587
  "(-1::real) / 3 = - 1 / 3"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   588
  "(-1::real) / -3 = 1 / 3"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   589
  "(x::real) / 1 = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   590
  "x > 0 \<longrightarrow> x / 3 < x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   591
  "x < 0 \<longrightarrow> x / 3 > x"
48069
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 47155
diff changeset
   592
  using [[z3_with_extensions]]
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   593
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   594
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   595
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   596
  "(3::real) * (x / 3) = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   597
  "(x * 3) / 3 = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   598
  "x > 0 \<longrightarrow> 2 * x / 3 < x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   599
  "x < 0 \<longrightarrow> 2 * x / 3 > x"
48069
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 47155
diff changeset
   600
  using [[z3_with_extensions]]
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   601
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   602
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   603
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   604
  "abs (x::real) \<ge> 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   605
  "(abs x = 0) = (x = 0)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   606
  "(x \<ge> 0) = (abs x = x)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   607
  "(x \<le> 0) = (abs x = -x)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   608
  "abs (abs x) = abs x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   609
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   610
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   611
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   612
  "min (x::real) y \<le> x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   613
  "min x y \<le> y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   614
  "z < x \<and> z < y \<longrightarrow> z < min x y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   615
  "min x y = min y x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   616
  "x \<ge> 0 \<longrightarrow> min x 0 = 0"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   617
  "min x y \<le> abs (x + y)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   618
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   619
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   620
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   621
  "max (x::real) y \<ge> x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   622
  "max x y \<ge> y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   623
  "z > x \<and> z > y \<longrightarrow> z > max x y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   624
  "max x y = max y x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   625
  "x \<ge> 0 \<longrightarrow> max x 0 = x"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   626
  "max x y \<ge> - abs x - abs y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   627
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   628
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   629
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   630
  "x \<le> (x::real)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   631
  "x \<le> y \<longrightarrow> 3 * x \<le> 3 * y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   632
  "x < y \<longrightarrow> 3 * x < 3 * y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   633
  "x < y \<longrightarrow> x \<le> y"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   634
  "\<not>(x < x)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   635
  "x \<le> y \<longrightarrow> y \<le> z \<longrightarrow> x \<le> z"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   636
  "x < y \<longrightarrow> y \<le> z \<longrightarrow> x \<le> z"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   637
  "x \<le> y \<longrightarrow> y < z \<longrightarrow> x \<le> z"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   638
  "x < y \<longrightarrow> y < z \<longrightarrow> x < z"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   639
  "x < y \<and> y < z \<longrightarrow> \<not>(z < x)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   640
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   641
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   642
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   643
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   644
section {* Datatypes, Records, and Typedefs *}
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   645
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   646
subsection {* Without support by the SMT solver *}
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   647
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   648
subsubsection {* Algebraic datatypes *}
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   649
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   650
lemma
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   651
  "x = fst (x, y)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   652
  "y = snd (x, y)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   653
  "((x, y) = (y, x)) = (x = y)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   654
  "((x, y) = (u, v)) = (x = u \<and> y = v)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   655
  "(fst (x, y, z) = fst (u, v, w)) = (x = u)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   656
  "(snd (x, y, z) = snd (u, v, w)) = (y = v \<and> z = w)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   657
  "(fst (snd (x, y, z)) = fst (snd (u, v, w))) = (y = v)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   658
  "(snd (snd (x, y, z)) = snd (snd (u, v, w))) = (z = w)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   659
  "(fst (x, y) = snd (x, y)) = (x = y)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   660
  "p1 = (x, y) \<and> p2 = (y, x) \<longrightarrow> fst p1 = snd p2"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   661
  "(fst (x, y) = snd (x, y)) = (x = y)"
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   662
  "(fst p = snd p) = (p = (snd p, fst p))"
41132
42384824b732 updated SMT certificates
boehmes
parents: 40681
diff changeset
   663
  using fst_conv snd_conv pair_collapse
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   664
  by smt+
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   665
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   666
lemma
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   667
  "[x] \<noteq> Nil"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   668
  "[x, y] \<noteq> Nil"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   669
  "x \<noteq> y \<longrightarrow> [x] \<noteq> [y]"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   670
  "hd (x # xs) = x"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   671
  "tl (x # xs) = xs"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   672
  "hd [x, y, z] = x"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   673
  "tl [x, y, z] = [y, z]"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   674
  "hd (tl [x, y, z]) = y"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   675
  "tl (tl [x, y, z]) = [z]"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   676
  using hd.simps tl.simps(2) list.simps
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   677
  by smt+
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   678
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   679
lemma
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   680
  "fst (hd [(a, b)]) = a"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   681
  "snd (hd [(a, b)]) = b"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   682
  using fst_conv snd_conv pair_collapse hd.simps tl.simps(2) list.simps
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   683
  by smt+
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   684
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   685
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   686
subsubsection {* Records *}
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   687
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   688
record point =
41427
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   689
  cx :: int
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   690
  cy :: int
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   691
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   692
record bw_point = point +
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   693
  black :: bool
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   694
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   695
lemma
41427
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   696
  "p1 = p2 \<longrightarrow> cx p1 = cx p2"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   697
  "p1 = p2 \<longrightarrow> cy p1 = cy p2"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   698
  "cx p1 \<noteq> cx p2 \<longrightarrow> p1 \<noteq> p2"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   699
  "cy p1 \<noteq> cy p2 \<longrightarrow> p1 \<noteq> p2"
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   700
  using point.simps
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   701
  by smt+
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   702
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   703
lemma
41427
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   704
  "cx \<lparr> cx = 3, cy = 4 \<rparr> = 3"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   705
  "cy \<lparr> cx = 3, cy = 4 \<rparr> = 4"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   706
  "cx \<lparr> cx = 3, cy = 4 \<rparr> \<noteq> cy \<lparr> cx = 3, cy = 4 \<rparr>"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   707
  "\<lparr> cx = 3, cy = 4 \<rparr> \<lparr> cx := 5 \<rparr> = \<lparr> cx = 5, cy = 4 \<rparr>"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   708
  "\<lparr> cx = 3, cy = 4 \<rparr> \<lparr> cy := 6 \<rparr> = \<lparr> cx = 3, cy = 6 \<rparr>"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   709
  "p = \<lparr> cx = 3, cy = 4 \<rparr> \<longrightarrow> p \<lparr> cx := 3 \<rparr> \<lparr> cy := 4 \<rparr> = p"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   710
  "p = \<lparr> cx = 3, cy = 4 \<rparr> \<longrightarrow> p \<lparr> cy := 4 \<rparr> \<lparr> cx := 3 \<rparr> = p"
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   711
  using point.simps
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   712
  using [[z3_options="AUTO_CONFIG=false"]]
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   713
  by smt+
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   714
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   715
lemma
41427
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   716
  "cy (p \<lparr> cx := a \<rparr>) = cy p"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   717
  "cx (p \<lparr> cy := a \<rparr>) = cx p"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   718
  "p \<lparr> cx := 3 \<rparr> \<lparr> cy := 4 \<rparr> = p \<lparr> cy := 4 \<rparr> \<lparr> cx := 3 \<rparr>"
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   719
  sorry
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   720
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   721
lemma
41427
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   722
  "p1 = p2 \<longrightarrow> cx p1 = cx p2"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   723
  "p1 = p2 \<longrightarrow> cy p1 = cy p2"
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   724
  "p1 = p2 \<longrightarrow> black p1 = black p2"
41427
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   725
  "cx p1 \<noteq> cx p2 \<longrightarrow> p1 \<noteq> p2"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   726
  "cy p1 \<noteq> cy p2 \<longrightarrow> p1 \<noteq> p2"
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   727
  "black p1 \<noteq> black p2 \<longrightarrow> p1 \<noteq> p2"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   728
  using point.simps bw_point.simps
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   729
  by smt+
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   730
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   731
lemma
41427
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   732
  "cx \<lparr> cx = 3, cy = 4, black = b \<rparr> = 3"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   733
  "cy \<lparr> cx = 3, cy = 4, black = b \<rparr> = 4"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   734
  "black \<lparr> cx = 3, cy = 4, black = b \<rparr> = b"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   735
  "cx \<lparr> cx = 3, cy = 4, black = b \<rparr> \<noteq> cy \<lparr> cx = 3, cy = 4, black = b \<rparr>"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   736
  "\<lparr> cx = 3, cy = 4, black = b \<rparr> \<lparr> cx := 5 \<rparr> = \<lparr> cx = 5, cy = 4, black = b \<rparr>"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   737
  "\<lparr> cx = 3, cy = 4, black = b \<rparr> \<lparr> cy := 6 \<rparr> = \<lparr> cx = 3, cy = 6, black = b \<rparr>"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   738
  "p = \<lparr> cx = 3, cy = 4, black = True \<rparr> \<longrightarrow>
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   739
     p \<lparr> cx := 3 \<rparr> \<lparr> cy := 4 \<rparr> \<lparr> black := True \<rparr> = p"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   740
  "p = \<lparr> cx = 3, cy = 4, black = True \<rparr> \<longrightarrow>
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   741
     p \<lparr> cy := 4 \<rparr> \<lparr> black := True \<rparr> \<lparr> cx := 3 \<rparr> = p"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   742
  "p = \<lparr> cx = 3, cy = 4, black = True \<rparr> \<longrightarrow>
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   743
     p \<lparr> black := True \<rparr> \<lparr> cx := 3 \<rparr> \<lparr> cy := 4 \<rparr> = p"
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   744
  using point.simps bw_point.simps
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   745
  using [[z3_options="AUTO_CONFIG=false"]]
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   746
  by smt+
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   747
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   748
lemma
41427
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   749
  "\<lparr> cx = 3, cy = 4, black = b \<rparr> \<lparr> black := w \<rparr> = \<lparr> cx = 3, cy = 4, black = w \<rparr>"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   750
  "\<lparr> cx = 3, cy = 4, black = True \<rparr> \<lparr> black := False \<rparr> =
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   751
     \<lparr> cx = 3, cy = 4, black = False \<rparr>"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   752
  "p \<lparr> cx := 3 \<rparr> \<lparr> cy := 4 \<rparr> \<lparr> black := True \<rparr> =
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   753
     p \<lparr> black := True \<rparr> \<lparr> cy := 4 \<rparr> \<lparr> cx := 3 \<rparr>"
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   754
  sorry
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   755
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   756
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   757
subsubsection {* Type definitions *}
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   758
45694
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 44925
diff changeset
   759
definition "three = {1, 2, 3::int}"
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 44925
diff changeset
   760
49834
b27bbb021df1 discontinued obsolete typedef (open) syntax;
wenzelm
parents: 48069
diff changeset
   761
typedef three = three
45694
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 44925
diff changeset
   762
  unfolding three_def by auto
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   763
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   764
definition n1 where "n1 = Abs_three 1"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   765
definition n2 where "n2 = Abs_three 2"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   766
definition n3 where "n3 = Abs_three 3"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   767
definition nplus where "nplus n m = Abs_three (Rep_three n + Rep_three m)"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   768
41427
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   769
lemma three_def': "(n \<in> three) = (n = 1 \<or> n = 2 \<or> n = 3)"
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   770
  by (auto simp add: three_def)
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   771
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   772
lemma
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   773
  "n1 = n1"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   774
  "n2 = n2"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   775
  "n1 \<noteq> n2"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   776
  "nplus n1 n1 = n2"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   777
  "nplus n1 n2 = n3"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   778
  using n1_def n2_def n3_def nplus_def
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   779
  using three_def' Rep_three Abs_three_inverse
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   780
  using [[z3_options="AUTO_CONFIG=false"]]
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   781
  by smt+
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   782
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   783
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   784
subsection {* With support by the SMT solver (but without proofs) *}
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   785
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   786
subsubsection {* Algebraic datatypes *}
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   787
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   788
lemma
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   789
  "x = fst (x, y)"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   790
  "y = snd (x, y)"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   791
  "((x, y) = (y, x)) = (x = y)"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   792
  "((x, y) = (u, v)) = (x = u \<and> y = v)"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   793
  "(fst (x, y, z) = fst (u, v, w)) = (x = u)"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   794
  "(snd (x, y, z) = snd (u, v, w)) = (y = v \<and> z = w)"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   795
  "(fst (snd (x, y, z)) = fst (snd (u, v, w))) = (y = v)"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   796
  "(snd (snd (x, y, z)) = snd (snd (u, v, w))) = (z = w)"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   797
  "(fst (x, y) = snd (x, y)) = (x = y)"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   798
  "p1 = (x, y) \<and> p2 = (y, x) \<longrightarrow> fst p1 = snd p2"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   799
  "(fst (x, y) = snd (x, y)) = (x = y)"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   800
  "(fst p = snd p) = (p = (snd p, fst p))"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   801
  using fst_conv snd_conv pair_collapse
48069
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 47155
diff changeset
   802
  using [[smt_datatypes, smt_oracle, z3_with_extensions]]
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   803
  by smt+
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   804
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   805
lemma
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   806
  "[x] \<noteq> Nil"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   807
  "[x, y] \<noteq> Nil"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   808
  "x \<noteq> y \<longrightarrow> [x] \<noteq> [y]"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   809
  "hd (x # xs) = x"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   810
  "tl (x # xs) = xs"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   811
  "hd [x, y, z] = x"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   812
  "tl [x, y, z] = [y, z]"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   813
  "hd (tl [x, y, z]) = y"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   814
  "tl (tl [x, y, z]) = [z]"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   815
  using hd.simps tl.simps(2)
48069
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 47155
diff changeset
   816
  using [[smt_datatypes, smt_oracle, z3_with_extensions]]
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   817
  by smt+
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   818
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   819
lemma
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   820
  "fst (hd [(a, b)]) = a"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   821
  "snd (hd [(a, b)]) = b"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   822
  using fst_conv snd_conv pair_collapse hd.simps tl.simps(2)
48069
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 47155
diff changeset
   823
  using [[smt_datatypes, smt_oracle, z3_with_extensions]]
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   824
  by smt+
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   825
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   826
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   827
subsubsection {* Records *}
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   828
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   829
lemma
41427
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   830
  "p1 = p2 \<longrightarrow> cx p1 = cx p2"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   831
  "p1 = p2 \<longrightarrow> cy p1 = cy p2"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   832
  "cx p1 \<noteq> cx p2 \<longrightarrow> p1 \<noteq> p2"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   833
  "cy p1 \<noteq> cy p2 \<longrightarrow> p1 \<noteq> p2"
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   834
  using point.simps
48069
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 47155
diff changeset
   835
  using [[smt_datatypes, smt_oracle, z3_with_extensions]]
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   836
  using [[z3_options="AUTO_CONFIG=false"]]
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   837
  by smt+
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   838
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   839
lemma
41427
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   840
  "cx \<lparr> cx = 3, cy = 4 \<rparr> = 3"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   841
  "cy \<lparr> cx = 3, cy = 4 \<rparr> = 4"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   842
  "cx \<lparr> cx = 3, cy = 4 \<rparr> \<noteq> cy \<lparr> cx = 3, cy = 4 \<rparr>"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   843
  "\<lparr> cx = 3, cy = 4 \<rparr> \<lparr> cx := 5 \<rparr> = \<lparr> cx = 5, cy = 4 \<rparr>"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   844
  "\<lparr> cx = 3, cy = 4 \<rparr> \<lparr> cy := 6 \<rparr> = \<lparr> cx = 3, cy = 6 \<rparr>"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   845
  "p = \<lparr> cx = 3, cy = 4 \<rparr> \<longrightarrow> p \<lparr> cx := 3 \<rparr> \<lparr> cy := 4 \<rparr> = p"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   846
  "p = \<lparr> cx = 3, cy = 4 \<rparr> \<longrightarrow> p \<lparr> cy := 4 \<rparr> \<lparr> cx := 3 \<rparr> = p"
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   847
  using point.simps
48069
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 47155
diff changeset
   848
  using [[smt_datatypes, smt_oracle, z3_with_extensions]]
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   849
  using [[z3_options="AUTO_CONFIG=false"]]
47111
a4476e55a241 reintroduced broken proofs and regenerated certificates
blanchet
parents: 47108
diff changeset
   850
  by smt+
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   851
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   852
lemma
41427
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   853
  "cy (p \<lparr> cx := a \<rparr>) = cy p"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   854
  "cx (p \<lparr> cy := a \<rparr>) = cx p"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   855
  "p \<lparr> cx := 3 \<rparr> \<lparr> cy := 4 \<rparr> = p \<lparr> cy := 4 \<rparr> \<lparr> cx := 3 \<rparr>"
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   856
  using point.simps
48069
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 47155
diff changeset
   857
  using [[smt_datatypes, smt_oracle, z3_with_extensions]]
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   858
  using [[z3_options="AUTO_CONFIG=false"]]
47111
a4476e55a241 reintroduced broken proofs and regenerated certificates
blanchet
parents: 47108
diff changeset
   859
  by smt+
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   860
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   861
lemma
41427
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   862
  "p1 = p2 \<longrightarrow> cx p1 = cx p2"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   863
  "p1 = p2 \<longrightarrow> cy p1 = cy p2"
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   864
  "p1 = p2 \<longrightarrow> black p1 = black p2"
41427
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   865
  "cx p1 \<noteq> cx p2 \<longrightarrow> p1 \<noteq> p2"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   866
  "cy p1 \<noteq> cy p2 \<longrightarrow> p1 \<noteq> p2"
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   867
  "black p1 \<noteq> black p2 \<longrightarrow> p1 \<noteq> p2"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   868
  using point.simps bw_point.simps
48069
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 47155
diff changeset
   869
  using [[smt_datatypes, smt_oracle, z3_with_extensions]]
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   870
  by smt+
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   871
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   872
lemma
41427
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   873
  "cx \<lparr> cx = 3, cy = 4, black = b \<rparr> = 3"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   874
  "cy \<lparr> cx = 3, cy = 4, black = b \<rparr> = 4"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   875
  "black \<lparr> cx = 3, cy = 4, black = b \<rparr> = b"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   876
  "cx \<lparr> cx = 3, cy = 4, black = b \<rparr> \<noteq> cy \<lparr> cx = 3, cy = 4, black = b \<rparr>"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   877
  "\<lparr> cx = 3, cy = 4, black = b \<rparr> \<lparr> cx := 5 \<rparr> = \<lparr> cx = 5, cy = 4, black = b \<rparr>"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   878
  "\<lparr> cx = 3, cy = 4, black = b \<rparr> \<lparr> cy := 6 \<rparr> = \<lparr> cx = 3, cy = 6, black = b \<rparr>"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   879
  "p = \<lparr> cx = 3, cy = 4, black = True \<rparr> \<longrightarrow>
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   880
     p \<lparr> cx := 3 \<rparr> \<lparr> cy := 4 \<rparr> \<lparr> black := True \<rparr> = p"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   881
  "p = \<lparr> cx = 3, cy = 4, black = True \<rparr> \<longrightarrow>
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   882
     p \<lparr> cy := 4 \<rparr> \<lparr> black := True \<rparr> \<lparr> cx := 3 \<rparr> = p"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   883
  "p = \<lparr> cx = 3, cy = 4, black = True \<rparr> \<longrightarrow>
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   884
     p \<lparr> black := True \<rparr> \<lparr> cx := 3 \<rparr> \<lparr> cy := 4 \<rparr> = p"
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   885
  using point.simps bw_point.simps
48069
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 47155
diff changeset
   886
  using [[smt_datatypes, smt_oracle, z3_with_extensions]]
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   887
  using [[z3_options="AUTO_CONFIG=false"]]
47111
a4476e55a241 reintroduced broken proofs and regenerated certificates
blanchet
parents: 47108
diff changeset
   888
  by smt+
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   889
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   890
lemma
41427
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   891
  "\<lparr> cx = 3, cy = 4, black = b \<rparr> \<lparr> black := w \<rparr> = \<lparr> cx = 3, cy = 4, black = w \<rparr>"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   892
  "\<lparr> cx = 3, cy = 4, black = True \<rparr> \<lparr> black := False \<rparr> =
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   893
     \<lparr> cx = 3, cy = 4, black = False \<rparr>"
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   894
  sorry
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   895
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   896
lemma
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   897
  "p \<lparr> cx := 3 \<rparr> \<lparr> cy := 4 \<rparr> \<lparr> black := True \<rparr> =
5fbad0390649 updated SMT certificates
boehmes
parents: 41426
diff changeset
   898
     p \<lparr> black := True \<rparr> \<lparr> cy := 4 \<rparr> \<lparr> cx := 3 \<rparr>"
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   899
  using point.simps bw_point.simps
48069
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 47155
diff changeset
   900
  using [[smt_datatypes, smt_oracle, z3_with_extensions]]
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   901
  using [[z3_options="AUTO_CONFIG=false"]]
47111
a4476e55a241 reintroduced broken proofs and regenerated certificates
blanchet
parents: 47108
diff changeset
   902
  by smt
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   903
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   904
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   905
subsubsection {* Type definitions *}
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   906
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   907
lemma
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   908
  "n1 = n1"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   909
  "n2 = n2"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   910
  "n1 \<noteq> n2"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   911
  "nplus n1 n1 = n2"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   912
  "nplus n1 n2 = n3"
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   913
  using n1_def n2_def n3_def nplus_def
48069
e9b2782c4f99 restricted Z3 by default to a fragment where proof reconstruction should not fail (for better integration with Sledgehammer) -- the full set of supported Z3 features can still be used by enabling the configuration option "z3_with_extensions"
boehmes
parents: 47155
diff changeset
   914
  using [[smt_datatypes, smt_oracle, z3_with_extensions]]
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   915
  using [[z3_options="AUTO_CONFIG=false"]]
47111
a4476e55a241 reintroduced broken proofs and regenerated certificates
blanchet
parents: 47108
diff changeset
   916
  by smt+
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41132
diff changeset
   917
37157
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   918
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   919
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   920
section {* Function updates *}
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   921
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   922
lemma
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   923
  "(f (i := v)) i = v"
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   924
  "i1 \<noteq> i2 \<longrightarrow> (f (i1 := v)) i2 = f i2"
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   925
  "i1 \<noteq> i2 \<longrightarrow> (f (i1 := v1, i2 := v2)) i1 = v1"
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   926
  "i1 \<noteq> i2 \<longrightarrow> (f (i1 := v1, i2 := v2)) i2 = v2"
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   927
  "i1 = i2 \<longrightarrow> (f (i1 := v1, i2 := v2)) i1 = v2"
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   928
  "i1 = i2 \<longrightarrow> (f (i1 := v1, i2 := v2)) i1 = v2"
47155
ade3fc826af3 dropped support for List.distinct in binding to SMT solvers: only few applications benefited from this support, and in some cases the smt method fails due to its support for List.distinct
boehmes
parents: 47152
diff changeset
   929
  "i1 \<noteq> i2 \<and>i1 \<noteq> i3 \<and>  i2 \<noteq> i3 \<longrightarrow> (f (i1 := v1, i2 := v2)) i3 = f i3"
41132
42384824b732 updated SMT certificates
boehmes
parents: 40681
diff changeset
   930
  using fun_upd_same fun_upd_apply
37157
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   931
  by smt+
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   932
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   933
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   934
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   935
section {* Sets *}
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   936
44925
1db6baa40b0e observe distinction between sets and predicates
boehmes
parents: 44753
diff changeset
   937
lemma Empty: "x \<notin> {}" by simp
1db6baa40b0e observe distinction between sets and predicates
boehmes
parents: 44753
diff changeset
   938
1db6baa40b0e observe distinction between sets and predicates
boehmes
parents: 44753
diff changeset
   939
lemmas smt_sets = Empty UNIV_I Un_iff Int_iff
37157
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   940
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   941
lemma
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   942
  "x \<notin> {}"
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   943
  "x \<in> UNIV"
44925
1db6baa40b0e observe distinction between sets and predicates
boehmes
parents: 44753
diff changeset
   944
  "x \<in> A \<union> B \<longleftrightarrow> x \<in> A \<or> x \<in> B"
1db6baa40b0e observe distinction between sets and predicates
boehmes
parents: 44753
diff changeset
   945
  "x \<in> P \<union> {} \<longleftrightarrow> x \<in> P"
37157
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   946
  "x \<in> P \<union> UNIV"
44925
1db6baa40b0e observe distinction between sets and predicates
boehmes
parents: 44753
diff changeset
   947
  "x \<in> P \<union> Q \<longleftrightarrow> x \<in> Q \<union> P"
1db6baa40b0e observe distinction between sets and predicates
boehmes
parents: 44753
diff changeset
   948
  "x \<in> P \<union> P \<longleftrightarrow> x \<in> P"
1db6baa40b0e observe distinction between sets and predicates
boehmes
parents: 44753
diff changeset
   949
  "x \<in> P \<union> (Q \<union> R) \<longleftrightarrow> x \<in> (P \<union> Q) \<union> R"
1db6baa40b0e observe distinction between sets and predicates
boehmes
parents: 44753
diff changeset
   950
  "x \<in> A \<inter> B \<longleftrightarrow> x \<in> A \<and> x \<in> B"
37157
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   951
  "x \<notin> P \<inter> {}"
44925
1db6baa40b0e observe distinction between sets and predicates
boehmes
parents: 44753
diff changeset
   952
  "x \<in> P \<inter> UNIV \<longleftrightarrow> x \<in> P"
1db6baa40b0e observe distinction between sets and predicates
boehmes
parents: 44753
diff changeset
   953
  "x \<in> P \<inter> Q \<longleftrightarrow> x \<in> Q \<inter> P"
1db6baa40b0e observe distinction between sets and predicates
boehmes
parents: 44753
diff changeset
   954
  "x \<in> P \<inter> P \<longleftrightarrow> x \<in> P"
1db6baa40b0e observe distinction between sets and predicates
boehmes
parents: 44753
diff changeset
   955
  "x \<in> P \<inter> (Q \<inter> R) \<longleftrightarrow> x \<in> (P \<inter> Q) \<inter> R"
1db6baa40b0e observe distinction between sets and predicates
boehmes
parents: 44753
diff changeset
   956
  "{x. x \<in> P} = {y. y \<in> P}"
37157
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   957
  by (smt smt_sets)+
86872cbae9e9 added function update examples and set examples
boehmes
parents: 37151
diff changeset
   958
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
diff changeset
   959
end