author | wenzelm |
Wed, 07 Aug 2019 09:28:32 +0200 | |
changeset 70477 | 90acc6ce5beb |
parent 69605 | a96320074298 |
permissions | -rw-r--r-- |
66614
1f1c5d85d232
moved Nunchaku to Main; the goal is to move Nitpick out in the next 1-2 years
blanchet
parents:
66163
diff
changeset
|
1 |
(* Title: HOL/Nunchaku.thy |
1f1c5d85d232
moved Nunchaku to Main; the goal is to move Nitpick out in the next 1-2 years
blanchet
parents:
66163
diff
changeset
|
2 |
Author: Jasmin Blanchette, VU Amsterdam |
1f1c5d85d232
moved Nunchaku to Main; the goal is to move Nitpick out in the next 1-2 years
blanchet
parents:
66163
diff
changeset
|
3 |
Copyright 2015, 2016, 2017 |
64389 | 4 |
|
5 |
Nunchaku: Yet another counterexample generator for Isabelle/HOL. |
|
6 |
||
66614
1f1c5d85d232
moved Nunchaku to Main; the goal is to move Nitpick out in the next 1-2 years
blanchet
parents:
66163
diff
changeset
|
7 |
Nunchaku relies on an external program of the same name. The sources are |
1f1c5d85d232
moved Nunchaku to Main; the goal is to move Nitpick out in the next 1-2 years
blanchet
parents:
66163
diff
changeset
|
8 |
available at |
64389 | 9 |
|
10 |
https://github.com/nunchaku-inria |
|
64407 | 11 |
|
64469
488d4e627238
added Nunchaku component and tuned Nunchaku integration accordingly
blanchet
parents:
64407
diff
changeset
|
12 |
The "$NUNCHAKU_HOME" environment variable must be set to the absolute path to |
488d4e627238
added Nunchaku component and tuned Nunchaku integration accordingly
blanchet
parents:
64407
diff
changeset
|
13 |
the directory containing the "nunchaku" executable. The Isabelle components |
66637 | 14 |
for CVC4, Kodkodi, and SMBC are necessary to use these backend solvers. |
64389 | 15 |
*) |
16 |
||
17 |
theory Nunchaku |
|
66614
1f1c5d85d232
moved Nunchaku to Main; the goal is to move Nitpick out in the next 1-2 years
blanchet
parents:
66163
diff
changeset
|
18 |
imports Nitpick |
64389 | 19 |
keywords |
20 |
"nunchaku" :: diag and |
|
21 |
"nunchaku_params" :: thy_decl |
|
22 |
begin |
|
23 |
||
24 |
consts unreachable :: 'a |
|
25 |
||
26 |
definition The_unsafe :: "('a \<Rightarrow> bool) \<Rightarrow> 'a" where |
|
27 |
"The_unsafe = The" |
|
28 |
||
29 |
definition rmember :: "'a set \<Rightarrow> 'a \<Rightarrow> bool" where |
|
30 |
"rmember A x \<longleftrightarrow> x \<in> A" |
|
31 |
||
69605 | 32 |
ML_file \<open>Tools/Nunchaku/nunchaku_util.ML\<close> |
33 |
ML_file \<open>Tools/Nunchaku/nunchaku_collect.ML\<close> |
|
34 |
ML_file \<open>Tools/Nunchaku/nunchaku_problem.ML\<close> |
|
35 |
ML_file \<open>Tools/Nunchaku/nunchaku_translate.ML\<close> |
|
36 |
ML_file \<open>Tools/Nunchaku/nunchaku_model.ML\<close> |
|
37 |
ML_file \<open>Tools/Nunchaku/nunchaku_reconstruct.ML\<close> |
|
38 |
ML_file \<open>Tools/Nunchaku/nunchaku_display.ML\<close> |
|
39 |
ML_file \<open>Tools/Nunchaku/nunchaku_tool.ML\<close> |
|
40 |
ML_file \<open>Tools/Nunchaku/nunchaku.ML\<close> |
|
41 |
ML_file \<open>Tools/Nunchaku/nunchaku_commands.ML\<close> |
|
64389 | 42 |
|
43 |
hide_const (open) unreachable The_unsafe rmember |
|
44 |
||
45 |
end |