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