author | wenzelm |
Wed, 28 Dec 2022 12:30:18 +0100 | |
changeset 76798 | 69d8d16c5612 |
parent 74641 | 6f801e1073fa |
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 |
||
63167 | 8 |
section \<open>Examples Featuring Minipick, the Minimalistic Version of Nitpick\<close> |
45035 | 9 |
|
10 |
theory Mini_Nits |
|
11 |
imports Main |
|
12 |
begin |
|
13 |
||
69605 | 14 |
ML_file \<open>minipick.ML\<close> |
48891 | 15 |
|
74641
6f801e1073fa
prefer "sat_solver = MiniSat", to make examples work uniformly on all platforms;
wenzelm
parents:
69605
diff
changeset
|
16 |
nitpick_params [verbose, sat_solver = MiniSat, max_threads = 1, total_consts = smart] |
45035 | 17 |
|
63167 | 18 |
ML \<open> |
69597 | 19 |
val check = Minipick.minipick \<^context> |
20 |
val expect = Minipick.minipick_expect \<^context> |
|
45062
9598cada31b3
first step towards extending Minipick with more translations
blanchet
parents:
45035
diff
changeset
|
21 |
val none = expect "none" |
9598cada31b3
first step towards extending Minipick with more translations
blanchet
parents:
45035
diff
changeset
|
22 |
val genuine = expect "genuine" |
9598cada31b3
first step towards extending Minipick with more translations
blanchet
parents:
45035
diff
changeset
|
23 |
val unknown = expect "unknown" |
63167 | 24 |
\<close> |
45035 | 25 |
|
69597 | 26 |
ML \<open>genuine 1 \<^prop>\<open>x = Not\<close>\<close> |
27 |
ML \<open>none 1 \<^prop>\<open>\<exists>x. x = Not\<close>\<close> |
|
28 |
ML \<open>none 1 \<^prop>\<open>\<not> False\<close>\<close> |
|
29 |
ML \<open>genuine 1 \<^prop>\<open>\<not> True\<close>\<close> |
|
30 |
ML \<open>none 1 \<^prop>\<open>\<not> \<not> b \<longleftrightarrow> b\<close>\<close> |
|
31 |
ML \<open>none 1 \<^prop>\<open>True\<close>\<close> |
|
32 |
ML \<open>genuine 1 \<^prop>\<open>False\<close>\<close> |
|
33 |
ML \<open>genuine 1 \<^prop>\<open>True \<longleftrightarrow> False\<close>\<close> |
|
34 |
ML \<open>none 1 \<^prop>\<open>True \<longleftrightarrow> \<not> False\<close>\<close> |
|
35 |
ML \<open>none 4 \<^prop>\<open>\<forall>x. x = x\<close>\<close> |
|
36 |
ML \<open>none 4 \<^prop>\<open>\<exists>x. x = x\<close>\<close> |
|
37 |
ML \<open>none 1 \<^prop>\<open>\<forall>x. x = y\<close>\<close> |
|
38 |
ML \<open>genuine 2 \<^prop>\<open>\<forall>x. x = y\<close>\<close> |
|
39 |
ML \<open>none 2 \<^prop>\<open>\<exists>x. x = y\<close>\<close> |
|
40 |
ML \<open>none 2 \<^prop>\<open>\<forall>x::'a \<times> 'a. x = x\<close>\<close> |
|
41 |
ML \<open>none 2 \<^prop>\<open>\<exists>x::'a \<times> 'a. x = y\<close>\<close> |
|
42 |
ML \<open>genuine 2 \<^prop>\<open>\<forall>x::'a \<times> 'a. x = y\<close>\<close> |
|
43 |
ML \<open>none 2 \<^prop>\<open>\<exists>x::'a \<times> 'a. x = y\<close>\<close> |
|
44 |
ML \<open>none 1 \<^prop>\<open>All = Ex\<close>\<close> |
|
45 |
ML \<open>genuine 2 \<^prop>\<open>All = Ex\<close>\<close> |
|
46 |
ML \<open>none 1 \<^prop>\<open>All P = Ex P\<close>\<close> |
|
47 |
ML \<open>genuine 2 \<^prop>\<open>All P = Ex P\<close>\<close> |
|
48 |
ML \<open>none 4 \<^prop>\<open>x = y \<longrightarrow> P x = P y\<close>\<close> |
|
49 |
ML \<open>none 4 \<^prop>\<open>(x::'a \<times> 'a) = y \<longrightarrow> P x = P y\<close>\<close> |
|
50 |
ML \<open>none 2 \<^prop>\<open>(x::'a \<times> 'a) = y \<longrightarrow> P x y = P y x\<close>\<close> |
|
51 |
ML \<open>none 4 \<^prop>\<open>\<exists>x::'a \<times> 'a. x = y \<longrightarrow> P x = P y\<close>\<close> |
|
52 |
ML \<open>none 2 \<^prop>\<open>(x::'a \<Rightarrow> 'a) = y \<longrightarrow> P x = P y\<close>\<close> |
|
53 |
ML \<open>none 2 \<^prop>\<open>\<exists>x::'a \<Rightarrow> 'a. x = y \<longrightarrow> P x = P y\<close>\<close> |
|
54 |
ML \<open>genuine 1 \<^prop>\<open>(=) X = Ex\<close>\<close> |
|
55 |
ML \<open>none 2 \<^prop>\<open>\<forall>x::'a \<Rightarrow> 'a. x = x\<close>\<close> |
|
56 |
ML \<open>none 1 \<^prop>\<open>x = y\<close>\<close> |
|
57 |
ML \<open>genuine 1 \<^prop>\<open>x \<longleftrightarrow> y\<close>\<close> |
|
58 |
ML \<open>genuine 2 \<^prop>\<open>x = y\<close>\<close> |
|
59 |
ML \<open>genuine 1 \<^prop>\<open>X \<subseteq> Y\<close>\<close> |
|
60 |
ML \<open>none 1 \<^prop>\<open>P \<and> Q \<longleftrightarrow> Q \<and> P\<close>\<close> |
|
61 |
ML \<open>none 1 \<^prop>\<open>P \<and> Q \<longrightarrow> P\<close>\<close> |
|
62 |
ML \<open>none 1 \<^prop>\<open>P \<or> Q \<longleftrightarrow> Q \<or> P\<close>\<close> |
|
63 |
ML \<open>genuine 1 \<^prop>\<open>P \<or> Q \<longrightarrow> P\<close>\<close> |
|
64 |
ML \<open>none 1 \<^prop>\<open>(P \<longrightarrow> Q) \<longleftrightarrow> (\<not> P \<or> Q)\<close>\<close> |
|
65 |
ML \<open>none 4 \<^prop>\<open>{a} = {a, a}\<close>\<close> |
|
66 |
ML \<open>genuine 2 \<^prop>\<open>{a} = {a, b}\<close>\<close> |
|
67 |
ML \<open>genuine 1 \<^prop>\<open>{a} \<noteq> {a, b}\<close>\<close> |
|
68 |
ML \<open>none 4 \<^prop>\<open>{}\<^sup>+ = {}\<close>\<close> |
|
69 |
ML \<open>none 4 \<^prop>\<open>UNIV\<^sup>+ = UNIV\<close>\<close> |
|
70 |
ML \<open>none 4 \<^prop>\<open>(UNIV :: ('a \<times> 'b) set) - {} = UNIV\<close>\<close> |
|
71 |
ML \<open>none 4 \<^prop>\<open>{} - (UNIV :: ('a \<times> 'b) set) = {}\<close>\<close> |
|
72 |
ML \<open>none 1 \<^prop>\<open>{(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}\<close>\<close> |
|
73 |
ML \<open>genuine 2 \<^prop>\<open>{(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}\<close>\<close> |
|
74 |
ML \<open>none 4 \<^prop>\<open>a \<noteq> c \<Longrightarrow> {(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}\<close>\<close> |
|
75 |
ML \<open>none 4 \<^prop>\<open>A \<union> B = {x. x \<in> A \<or> x \<in> B}\<close>\<close> |
|
76 |
ML \<open>none 4 \<^prop>\<open>A \<inter> B = {x. x \<in> A \<and> x \<in> B}\<close>\<close> |
|
77 |
ML \<open>none 4 \<^prop>\<open>A - B = (\<lambda>x. A x \<and> \<not> B x)\<close>\<close> |
|
78 |
ML \<open>none 4 \<^prop>\<open>\<exists>a b. (a, b) = (b, a)\<close>\<close> |
|
79 |
ML \<open>genuine 2 \<^prop>\<open>(a, b) = (b, a)\<close>\<close> |
|
80 |
ML \<open>genuine 2 \<^prop>\<open>(a, b) \<noteq> (b, a)\<close>\<close> |
|
81 |
ML \<open>none 4 \<^prop>\<open>\<exists>a b::'a \<times> 'a. (a, b) = (b, a)\<close>\<close> |
|
82 |
ML \<open>genuine 2 \<^prop>\<open>(a::'a \<times> 'a, b) = (b, a)\<close>\<close> |
|
83 |
ML \<open>none 4 \<^prop>\<open>\<exists>a b::'a \<times> 'a \<times> 'a. (a, b) = (b, a)\<close>\<close> |
|
84 |
ML \<open>genuine 2 \<^prop>\<open>(a::'a \<times> 'a \<times> 'a, b) \<noteq> (b, a)\<close>\<close> |
|
85 |
ML \<open>none 4 \<^prop>\<open>\<exists>a b::'a \<Rightarrow> 'a. (a, b) = (b, a)\<close>\<close> |
|
86 |
ML \<open>genuine 1 \<^prop>\<open>(a::'a \<Rightarrow> 'a, b) \<noteq> (b, a)\<close>\<close> |
|
87 |
ML \<open>none 4 \<^prop>\<open>fst (a, b) = a\<close>\<close> |
|
88 |
ML \<open>none 1 \<^prop>\<open>fst (a, b) = b\<close>\<close> |
|
89 |
ML \<open>genuine 2 \<^prop>\<open>fst (a, b) = b\<close>\<close> |
|
90 |
ML \<open>genuine 2 \<^prop>\<open>fst (a, b) \<noteq> b\<close>\<close> |
|
91 |
ML \<open>genuine 2 \<^prop>\<open>f ((x, z), y) = (x, z)\<close>\<close> |
|
92 |
ML \<open>none 2 \<^prop>\<open>(\<forall>x. f x = fst x) \<longrightarrow> f ((x, z), y) = (x, z)\<close>\<close> |
|
93 |
ML \<open>none 4 \<^prop>\<open>snd (a, b) = b\<close>\<close> |
|
94 |
ML \<open>none 1 \<^prop>\<open>snd (a, b) = a\<close>\<close> |
|
95 |
ML \<open>genuine 2 \<^prop>\<open>snd (a, b) = a\<close>\<close> |
|
96 |
ML \<open>genuine 2 \<^prop>\<open>snd (a, b) \<noteq> a\<close>\<close> |
|
97 |
ML \<open>genuine 1 \<^prop>\<open>P\<close>\<close> |
|
98 |
ML \<open>genuine 1 \<^prop>\<open>(\<lambda>x. P) a\<close>\<close> |
|
99 |
ML \<open>genuine 1 \<^prop>\<open>(\<lambda>x y z. P y x z) a b c\<close>\<close> |
|
100 |
ML \<open>none 4 \<^prop>\<open>\<exists>f. f = (\<lambda>x. x) \<and> f y = y\<close>\<close> |
|
101 |
ML \<open>genuine 1 \<^prop>\<open>\<exists>f. f p \<noteq> p \<and> (\<forall>a b. f (a, b) = (a, b))\<close>\<close> |
|
102 |
ML \<open>none 2 \<^prop>\<open>\<exists>f. \<forall>a b. f (a, b) = (a, b)\<close>\<close> |
|
103 |
ML \<open>none 3 \<^prop>\<open>f = (\<lambda>a b. (b, a)) \<longrightarrow> f x y = (y, x)\<close>\<close> |
|
104 |
ML \<open>genuine 2 \<^prop>\<open>f = (\<lambda>a b. (b, a)) \<longrightarrow> f x y = (x, y)\<close>\<close> |
|
105 |
ML \<open>none 4 \<^prop>\<open>f = (\<lambda>x. f x)\<close>\<close> |
|
106 |
ML \<open>none 4 \<^prop>\<open>f = (\<lambda>x. f x::'a \<Rightarrow> bool)\<close>\<close> |
|
107 |
ML \<open>none 4 \<^prop>\<open>f = (\<lambda>x y. f x y)\<close>\<close> |
|
108 |
ML \<open>none 4 \<^prop>\<open>f = (\<lambda>x y. f x y::'a \<Rightarrow> bool)\<close>\<close> |
|
45035 | 109 |
|
110 |
end |