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