author | wenzelm |
Thu, 01 Dec 2011 20:54:48 +0100 | |
changeset 45717 | b4e7b9968e60 |
parent 45062 | 9598cada31b3 |
child 45970 | b6d0cff57d96 |
permissions | -rw-r--r-- |
45035 | 1 |
(* Title: HOL/Nitpick_Examples/Mini_Nits.thy |
2 |
Author: Jasmin Blanchette, TU Muenchen |
|
3 |
Copyright 2009-2011 |
|
4 |
||
5 |
Examples featuring Minipick, the minimalistic version of Nitpick. |
|
6 |
*) |
|
7 |
||
8 |
header {* Examples Featuring Minipick, the Minimalistic Version of Nitpick *} |
|
9 |
||
10 |
theory Mini_Nits |
|
11 |
imports Main |
|
12 |
uses "minipick.ML" |
|
13 |
begin |
|
14 |
||
45062
9598cada31b3
first step towards extending Minipick with more translations
blanchet
parents:
45035
diff
changeset
|
15 |
nitpick_params [verbose, sat_solver = MiniSat_JNI, max_threads = 1] |
45035 | 16 |
|
45062
9598cada31b3
first step towards extending Minipick with more translations
blanchet
parents:
45035
diff
changeset
|
17 |
nitpick_params [total_consts = smart] |
45035 | 18 |
|
45062
9598cada31b3
first step towards extending Minipick with more translations
blanchet
parents:
45035
diff
changeset
|
19 |
ML {* |
9598cada31b3
first step towards extending Minipick with more translations
blanchet
parents:
45035
diff
changeset
|
20 |
val check = Minipick.minipick @{context} |
9598cada31b3
first step towards extending Minipick with more translations
blanchet
parents:
45035
diff
changeset
|
21 |
val expect = Minipick.minipick_expect @{context} |
9598cada31b3
first step towards extending Minipick with more translations
blanchet
parents:
45035
diff
changeset
|
22 |
val none = expect "none" |
9598cada31b3
first step towards extending Minipick with more translations
blanchet
parents:
45035
diff
changeset
|
23 |
val genuine = expect "genuine" |
9598cada31b3
first step towards extending Minipick with more translations
blanchet
parents:
45035
diff
changeset
|
24 |
val unknown = expect "unknown" |
45035 | 25 |
*} |
26 |
||
27 |
ML {* genuine 1 @{prop "x = Not"} *} |
|
28 |
ML {* none 1 @{prop "\<exists>x. x = Not"} *} |
|
29 |
ML {* none 1 @{prop "\<not> False"} *} |
|
30 |
ML {* genuine 1 @{prop "\<not> True"} *} |
|
31 |
ML {* none 1 @{prop "\<not> \<not> b \<longleftrightarrow> b"} *} |
|
32 |
ML {* none 1 @{prop True} *} |
|
33 |
ML {* genuine 1 @{prop False} *} |
|
34 |
ML {* genuine 1 @{prop "True \<longleftrightarrow> False"} *} |
|
35 |
ML {* none 1 @{prop "True \<longleftrightarrow> \<not> False"} *} |
|
36 |
ML {* none 5 @{prop "\<forall>x. x = x"} *} |
|
37 |
ML {* none 5 @{prop "\<exists>x. x = x"} *} |
|
38 |
ML {* none 1 @{prop "\<forall>x. x = y"} *} |
|
39 |
ML {* genuine 2 @{prop "\<forall>x. x = y"} *} |
|
40 |
ML {* none 2 @{prop "\<exists>x. x = y"} *} |
|
41 |
ML {* none 2 @{prop "\<forall>x\<Colon>'a \<times> 'a. x = x"} *} |
|
42 |
ML {* none 2 @{prop "\<exists>x\<Colon>'a \<times> 'a. x = y"} *} |
|
43 |
ML {* genuine 2 @{prop "\<forall>x\<Colon>'a \<times> 'a. x = y"} *} |
|
44 |
ML {* none 2 @{prop "\<exists>x\<Colon>'a \<times> 'a. x = y"} *} |
|
45 |
ML {* none 1 @{prop "All = Ex"} *} |
|
46 |
ML {* genuine 2 @{prop "All = Ex"} *} |
|
47 |
ML {* none 1 @{prop "All P = Ex P"} *} |
|
48 |
ML {* genuine 2 @{prop "All P = Ex P"} *} |
|
49 |
ML {* none 5 @{prop "x = y \<longrightarrow> P x = P y"} *} |
|
50 |
ML {* none 5 @{prop "(x\<Colon>'a \<times> 'a) = y \<longrightarrow> P x = P y"} *} |
|
51 |
ML {* none 2 @{prop "(x\<Colon>'a \<times> 'a) = y \<longrightarrow> P x y = P y x"} *} |
|
52 |
ML {* none 5 @{prop "\<exists>x\<Colon>'a \<times> 'a. x = y \<longrightarrow> P x = P y"} *} |
|
53 |
ML {* none 2 @{prop "(x\<Colon>'a \<Rightarrow> 'a) = y \<longrightarrow> P x = P y"} *} |
|
54 |
ML {* none 2 @{prop "\<exists>x\<Colon>'a \<Rightarrow> 'a. x = y \<longrightarrow> P x = P y"} *} |
|
55 |
ML {* genuine 1 @{prop "(op =) X = Ex"} *} |
|
56 |
ML {* none 2 @{prop "\<forall>x::'a \<Rightarrow> 'a. x = x"} *} |
|
57 |
ML {* none 1 @{prop "x = y"} *} |
|
58 |
ML {* genuine 1 @{prop "x \<longleftrightarrow> y"} *} |
|
59 |
ML {* genuine 2 @{prop "x = y"} *} |
|
60 |
ML {* genuine 1 @{prop "X \<subseteq> Y"} *} |
|
61 |
ML {* none 1 @{prop "P \<and> Q \<longleftrightarrow> Q \<and> P"} *} |
|
62 |
ML {* none 1 @{prop "P \<and> Q \<longrightarrow> P"} *} |
|
63 |
ML {* none 1 @{prop "P \<or> Q \<longleftrightarrow> Q \<or> P"} *} |
|
64 |
ML {* genuine 1 @{prop "P \<or> Q \<longrightarrow> P"} *} |
|
65 |
ML {* none 1 @{prop "(P \<longrightarrow> Q) \<longleftrightarrow> (\<not> P \<or> Q)"} *} |
|
66 |
ML {* none 5 @{prop "{a} = {a, a}"} *} |
|
67 |
ML {* genuine 2 @{prop "{a} = {a, b}"} *} |
|
68 |
ML {* genuine 1 @{prop "{a} \<noteq> {a, b}"} *} |
|
69 |
ML {* none 5 @{prop "{}\<^sup>+ = {}"} *} |
|
45062
9598cada31b3
first step towards extending Minipick with more translations
blanchet
parents:
45035
diff
changeset
|
70 |
ML {* none 5 @{prop "UNIV\<^sup>+ = UNIV"} *} |
9598cada31b3
first step towards extending Minipick with more translations
blanchet
parents:
45035
diff
changeset
|
71 |
ML {* none 5 @{prop "(UNIV\<Colon>'a \<times> 'b \<Rightarrow> bool) - {} = UNIV"} *} |
9598cada31b3
first step towards extending Minipick with more translations
blanchet
parents:
45035
diff
changeset
|
72 |
ML {* none 5 @{prop "{} - (UNIV\<Colon>'a \<times> 'b \<Rightarrow> bool) = {}"} *} |
45035 | 73 |
ML {* none 1 @{prop "{(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}"} *} |
74 |
ML {* genuine 2 @{prop "{(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}"} *} |
|
75 |
ML {* none 5 @{prop "a \<noteq> c \<Longrightarrow> {(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}"} *} |
|
76 |
ML {* none 5 @{prop "A \<union> B = (\<lambda>x. A x \<or> B x)"} *} |
|
77 |
ML {* none 5 @{prop "A \<inter> B = (\<lambda>x. A x \<and> B x)"} *} |
|
78 |
ML {* none 5 @{prop "A - B = (\<lambda>x. A x \<and> \<not> B x)"} *} |
|
79 |
ML {* none 5 @{prop "\<exists>a b. (a, b) = (b, a)"} *} |
|
80 |
ML {* genuine 2 @{prop "(a, b) = (b, a)"} *} |
|
81 |
ML {* genuine 2 @{prop "(a, b) \<noteq> (b, a)"} *} |
|
82 |
ML {* none 5 @{prop "\<exists>a b\<Colon>'a \<times> 'a. (a, b) = (b, a)"} *} |
|
83 |
ML {* genuine 2 @{prop "(a\<Colon>'a \<times> 'a, b) = (b, a)"} *} |
|
84 |
ML {* none 5 @{prop "\<exists>a b\<Colon>'a \<times> 'a \<times> 'a. (a, b) = (b, a)"} *} |
|
85 |
ML {* genuine 2 @{prop "(a\<Colon>'a \<times> 'a \<times> 'a, b) \<noteq> (b, a)"} *} |
|
86 |
ML {* none 5 @{prop "\<exists>a b\<Colon>'a \<Rightarrow> 'a. (a, b) = (b, a)"} *} |
|
87 |
ML {* genuine 1 @{prop "(a\<Colon>'a \<Rightarrow> 'a, b) \<noteq> (b, a)"} *} |
|
88 |
ML {* none 5 @{prop "fst (a, b) = a"} *} |
|
89 |
ML {* none 1 @{prop "fst (a, b) = b"} *} |
|
90 |
ML {* genuine 2 @{prop "fst (a, b) = b"} *} |
|
91 |
ML {* genuine 2 @{prop "fst (a, b) \<noteq> b"} *} |
|
45062
9598cada31b3
first step towards extending Minipick with more translations
blanchet
parents:
45035
diff
changeset
|
92 |
ML {* genuine 2 @{prop "f ((x, z), y) = (x, z)"} *} |
9598cada31b3
first step towards extending Minipick with more translations
blanchet
parents:
45035
diff
changeset
|
93 |
ML {* none 2 @{prop "(ALL x. f x = fst x) \<longrightarrow> f ((x, z), y) = (x, z)"} *} |
45035 | 94 |
ML {* none 5 @{prop "snd (a, b) = b"} *} |
95 |
ML {* none 1 @{prop "snd (a, b) = a"} *} |
|
96 |
ML {* genuine 2 @{prop "snd (a, b) = a"} *} |
|
97 |
ML {* genuine 2 @{prop "snd (a, b) \<noteq> a"} *} |
|
98 |
ML {* genuine 1 @{prop P} *} |
|
99 |
ML {* genuine 1 @{prop "(\<lambda>x. P) a"} *} |
|
100 |
ML {* genuine 1 @{prop "(\<lambda>x y z. P y x z) a b c"} *} |
|
101 |
ML {* none 5 @{prop "\<exists>f. f = (\<lambda>x. x) \<and> f y = y"} *} |
|
102 |
ML {* genuine 1 @{prop "\<exists>f. f p \<noteq> p \<and> (\<forall>a b. f (a, b) = (a, b))"} *} |
|
103 |
ML {* none 2 @{prop "\<exists>f. \<forall>a b. f (a, b) = (a, b)"} *} |
|
104 |
ML {* none 3 @{prop "f = (\<lambda>a b. (b, a)) \<longrightarrow> f x y = (y, x)"} *} |
|
105 |
ML {* genuine 2 @{prop "f = (\<lambda>a b. (b, a)) \<longrightarrow> f x y = (x, y)"} *} |
|
45062
9598cada31b3
first step towards extending Minipick with more translations
blanchet
parents:
45035
diff
changeset
|
106 |
ML {* none 5 @{prop "f = (\<lambda>x. f x)"} *} |
9598cada31b3
first step towards extending Minipick with more translations
blanchet
parents:
45035
diff
changeset
|
107 |
ML {* none 5 @{prop "f = (\<lambda>x. f x\<Colon>'a \<Rightarrow> bool)"} *} |
9598cada31b3
first step towards extending Minipick with more translations
blanchet
parents:
45035
diff
changeset
|
108 |
ML {* none 5 @{prop "f = (\<lambda>x y. f x y)"} *} |
9598cada31b3
first step towards extending Minipick with more translations
blanchet
parents:
45035
diff
changeset
|
109 |
ML {* none 5 @{prop "f = (\<lambda>x y. f x y\<Colon>'a \<Rightarrow> bool)"} *} |
45035 | 110 |
|
111 |
end |