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