| author | haftmann | 
| Sat, 19 Oct 2019 09:15:41 +0000 | |
| changeset 70903 | c550368a4e29 | 
| parent 70008 | 7aaebfcf3134 | 
| child 74399 | a1d33d1bfb6d | 
| permissions | -rw-r--r-- | 
| 33197 | 1  | 
(* Title: HOL/Nitpick_Examples/Mono_Nits.thy  | 
2  | 
Author: Jasmin Blanchette, TU Muenchen  | 
|
| 45035 | 3  | 
Copyright 2009-2011  | 
| 33197 | 4  | 
|
5  | 
Examples featuring Nitpick's monotonicity check.  | 
|
6  | 
*)  | 
|
7  | 
||
| 63167 | 8  | 
section \<open>Examples Featuring Nitpick's Monotonicity Check\<close>  | 
| 33197 | 9  | 
|
10  | 
theory Mono_Nits  | 
|
| 
46082
 
1c436a920730
reintroduced failing examples now that they work again, after reintroduction of "set"
 
blanchet 
parents: 
45972 
diff
changeset
 | 
11  | 
imports Main  | 
| 
 
1c436a920730
reintroduced failing examples now that they work again, after reintroduction of "set"
 
blanchet 
parents: 
45972 
diff
changeset
 | 
12  | 
(* "~/afp/thys/DPT-SAT-Solver/DPT_SAT_Solver" *)  | 
| 
 
1c436a920730
reintroduced failing examples now that they work again, after reintroduction of "set"
 
blanchet 
parents: 
45972 
diff
changeset
 | 
13  | 
(* "~/afp/thys/AVL-Trees/AVL2" "~/afp/thys/Huffman/Huffman" *)  | 
| 33197 | 14  | 
begin  | 
15  | 
||
| 63167 | 16  | 
ML \<open>  | 
| 
41008
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
17  | 
open Nitpick_Util  | 
| 
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
18  | 
open Nitpick_HOL  | 
| 41010 | 19  | 
open Nitpick_Preproc  | 
| 
41008
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
20  | 
|
| 40992 | 21  | 
exception BUG  | 
| 33197 | 22  | 
|
| 69597 | 23  | 
val thy = \<^theory>  | 
24  | 
val ctxt = \<^context>  | 
|
| 35339 | 25  | 
val subst = []  | 
| 54833 | 26  | 
val tac_timeout = seconds 1.0  | 
| 55888 | 27  | 
val case_names = case_const_names ctxt  | 
| 
42415
 
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
 
blanchet 
parents: 
41876 
diff
changeset
 | 
28  | 
val defs = all_defs_of thy subst  | 
| 
 
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
 
blanchet 
parents: 
41876 
diff
changeset
 | 
29  | 
val nondefs = all_nondefs_of ctxt subst  | 
| 
41791
 
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
 
blanchet 
parents: 
41523 
diff
changeset
 | 
30  | 
val def_tables = const_def_tables ctxt subst defs  | 
| 
42415
 
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
 
blanchet 
parents: 
41876 
diff
changeset
 | 
31  | 
val nondef_table = const_nondef_table nondefs  | 
| 
41008
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
32  | 
val simp_table = Unsynchronized.ref (const_simp_table ctxt subst)  | 
| 
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
33  | 
val psimp_table = const_psimp_table ctxt subst  | 
| 
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
34  | 
val choice_spec_table = const_choice_spec_table ctxt subst  | 
| 
41791
 
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
 
blanchet 
parents: 
41523 
diff
changeset
 | 
35  | 
val intro_table = inductive_intro_table ctxt subst def_tables  | 
| 
41008
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
36  | 
val ground_thm_table = ground_theorem_table thy  | 
| 
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
37  | 
val ersatz_table = ersatz_table ctxt  | 
| 41410 | 38  | 
val hol_ctxt as {thy, ...} : hol_context =
 | 
| 55888 | 39  | 
  {thy = thy, ctxt = ctxt, max_bisim_depth = ~1, boxes = [], wfs = [],
 | 
40  | 
user_axioms = NONE, debug = false, whacks = [], binary_ints = SOME false,  | 
|
41  | 
destroy_constrs = true, specialize = false, star_linear_preds = false,  | 
|
42  | 
total_consts = NONE, needs = NONE, tac_timeout = tac_timeout, evals = [],  | 
|
43  | 
case_names = case_names, def_tables = def_tables,  | 
|
44  | 
nondef_table = nondef_table, nondefs = nondefs, simp_table = simp_table,  | 
|
45  | 
psimp_table = psimp_table, choice_spec_table = choice_spec_table,  | 
|
46  | 
intro_table = intro_table, ground_thm_table = ground_thm_table,  | 
|
47  | 
ersatz_table = ersatz_table, skolems = Unsynchronized.ref [],  | 
|
48  | 
special_funs = Unsynchronized.ref [], unrolled_preds = Unsynchronized.ref [],  | 
|
49  | 
wf_cache = Unsynchronized.ref [], constr_cache = Unsynchronized.ref []}  | 
|
| 
41008
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
50  | 
val binarize = false  | 
| 
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
51  | 
|
| 
35812
 
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
 
blanchet 
parents: 
35807 
diff
changeset
 | 
52  | 
fun is_mono t =  | 
| 69597 | 53  | 
Nitpick_Mono.formulas_monotonic hol_ctxt binarize \<^typ>\<open>'a\<close> ([t], [])  | 
| 
41008
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
54  | 
|
| 33197 | 55  | 
fun is_const t =  | 
| 
46082
 
1c436a920730
reintroduced failing examples now that they work again, after reintroduction of "set"
 
blanchet 
parents: 
45972 
diff
changeset
 | 
56  | 
let val T = fastype_of t in  | 
| 69597 | 57  | 
    Logic.mk_implies (Logic.mk_equals (Free ("dummyP", T), t), \<^const>\<open>False\<close>)
 | 
| 
46082
 
1c436a920730
reintroduced failing examples now that they work again, after reintroduction of "set"
 
blanchet 
parents: 
45972 
diff
changeset
 | 
58  | 
|> is_mono  | 
| 33197 | 59  | 
end  | 
| 
41008
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
60  | 
|
| 40992 | 61  | 
fun mono t = is_mono t orelse raise BUG  | 
62  | 
fun nonmono t = not (is_mono t) orelse raise BUG  | 
|
63  | 
fun const t = is_const t orelse raise BUG  | 
|
64  | 
fun nonconst t = not (is_const t) orelse raise BUG  | 
|
| 63167 | 65  | 
\<close>  | 
| 33197 | 66  | 
|
| 63167 | 67  | 
ML \<open>Nitpick_Mono.trace := false\<close>  | 
| 
41006
 
000ca46429cd
added examples to exercise new monotonicity code
 
blanchet 
parents: 
40992 
diff
changeset
 | 
68  | 
|
| 69597 | 69  | 
ML_val \<open>const \<^term>\<open>A::('a\<Rightarrow>'b)\<close>\<close>
 | 
70  | 
ML_val \<open>const \<^term>\<open>(A::'a set) = A\<close>\<close>  | 
|
71  | 
ML_val \<open>const \<^term>\<open>(A::'a set set) = A\<close>\<close>  | 
|
72  | 
ML_val \<open>const \<^term>\<open>(\<lambda>x::'a set. a \<in> x)\<close>\<close>  | 
|
73  | 
ML_val \<open>const \<^term>\<open>{{a::'a}} = C\<close>\<close>
 | 
|
74  | 
ML_val \<open>const \<^term>\<open>{f::'a\<Rightarrow>nat} = {g::'a\<Rightarrow>nat}\<close>\<close>
 | 
|
75  | 
ML_val \<open>const \<^term>\<open>A \<union> (B::'a set)\<close>\<close>  | 
|
76  | 
ML_val \<open>const \<^term>\<open>\<lambda>A B x::'a. A x \<or> B x\<close>\<close>  | 
|
77  | 
ML_val \<open>const \<^term>\<open>P (a::'a)\<close>\<close>  | 
|
78  | 
ML_val \<open>const \<^term>\<open>\<lambda>a::'a. b (c (d::'a)) (e::'a) (f::'a)\<close>\<close>  | 
|
79  | 
ML_val \<open>const \<^term>\<open>\<forall>A::'a set. a \<in> A\<close>\<close>  | 
|
80  | 
ML_val \<open>const \<^term>\<open>\<forall>A::'a set. P A\<close>\<close>  | 
|
81  | 
ML_val \<open>const \<^term>\<open>P \<or> Q\<close>\<close>  | 
|
82  | 
ML_val \<open>const \<^term>\<open>A \<union> B = (C::'a set)\<close>\<close>  | 
|
83  | 
ML_val \<open>const \<^term>\<open>(\<lambda>A B x::'a. A x \<or> B x) A B = C\<close>\<close>  | 
|
84  | 
ML_val \<open>const \<^term>\<open>(if P then (A::'a set) else B) = C\<close>\<close>  | 
|
85  | 
ML_val \<open>const \<^term>\<open>let A = (C::'a set) in A \<union> B\<close>\<close>  | 
|
86  | 
ML_val \<open>const \<^term>\<open>THE x::'b. P x\<close>\<close>  | 
|
87  | 
ML_val \<open>const \<^term>\<open>(\<lambda>x::'a. False)\<close>\<close>  | 
|
88  | 
ML_val \<open>const \<^term>\<open>(\<lambda>x::'a. True)\<close>\<close>  | 
|
89  | 
ML_val \<open>const \<^term>\<open>(\<lambda>x::'a. False) = (\<lambda>x::'a. False)\<close>\<close>  | 
|
90  | 
ML_val \<open>const \<^term>\<open>(\<lambda>x::'a. True) = (\<lambda>x::'a. True)\<close>\<close>  | 
|
91  | 
ML_val \<open>const \<^term>\<open>Let (a::'a) A\<close>\<close>  | 
|
92  | 
ML_val \<open>const \<^term>\<open>A (a::'a)\<close>\<close>  | 
|
93  | 
ML_val \<open>const \<^term>\<open>insert (a::'a) A = B\<close>\<close>  | 
|
94  | 
ML_val \<open>const \<^term>\<open>- (A::'a set)\<close>\<close>  | 
|
95  | 
ML_val \<open>const \<^term>\<open>finite (A::'a set)\<close>\<close>  | 
|
96  | 
ML_val \<open>const \<^term>\<open>\<not> finite (A::'a set)\<close>\<close>  | 
|
97  | 
ML_val \<open>const \<^term>\<open>finite (A::'a set set)\<close>\<close>  | 
|
98  | 
ML_val \<open>const \<^term>\<open>\<lambda>a::'a. A a \<and> \<not> B a\<close>\<close>  | 
|
99  | 
ML_val \<open>const \<^term>\<open>A < (B::'a set)\<close>\<close>  | 
|
100  | 
ML_val \<open>const \<^term>\<open>A \<le> (B::'a set)\<close>\<close>  | 
|
101  | 
ML_val \<open>const \<^term>\<open>[a::'a]\<close>\<close>  | 
|
102  | 
ML_val \<open>const \<^term>\<open>[a::'a set]\<close>\<close>  | 
|
103  | 
ML_val \<open>const \<^term>\<open>[A \<union> (B::'a set)]\<close>\<close>  | 
|
104  | 
ML_val \<open>const \<^term>\<open>[A \<union> (B::'a set)] = [C]\<close>\<close>  | 
|
105  | 
ML_val \<open>const \<^term>\<open>{(\<lambda>x::'a. x = a)} = C\<close>\<close>
 | 
|
106  | 
ML_val \<open>const \<^term>\<open>(\<lambda>a::'a. \<not> A a) = B\<close>\<close>  | 
|
107  | 
ML_val \<open>const \<^prop>\<open>\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> f a \<and> g a \<longrightarrow> \<not> f a\<close>\<close>  | 
|
108  | 
ML_val \<open>const \<^term>\<open>\<lambda>A B x::'a. A x \<and> B x \<and> A = B\<close>\<close>  | 
|
109  | 
ML_val \<open>const \<^term>\<open>p = (\<lambda>(x::'a) (y::'a). P x \<or> \<not> Q y)\<close>\<close>  | 
|
110  | 
ML_val \<open>const \<^term>\<open>p = (\<lambda>(x::'a) (y::'a). p x y :: bool)\<close>\<close>  | 
|
111  | 
ML_val \<open>const \<^term>\<open>p = (\<lambda>A B x. A x \<and> \<not> B x) (\<lambda>x. True) (\<lambda>y. x \<noteq> y)\<close>\<close>  | 
|
112  | 
ML_val \<open>const \<^term>\<open>p = (\<lambda>y. x \<noteq> y)\<close>\<close>  | 
|
113  | 
ML_val \<open>const \<^term>\<open>(\<lambda>x. (p::'a\<Rightarrow>bool\<Rightarrow>bool) x False)\<close>\<close>  | 
|
114  | 
ML_val \<open>const \<^term>\<open>(\<lambda>x y. (p::'a\<Rightarrow>'a\<Rightarrow>bool\<Rightarrow>bool) x y False)\<close>\<close>  | 
|
115  | 
ML_val \<open>const \<^term>\<open>f = (\<lambda>x::'a. P x \<longrightarrow> Q x)\<close>\<close>  | 
|
116  | 
ML_val \<open>const \<^term>\<open>\<forall>a::'a. P a\<close>\<close>  | 
|
| 33197 | 117  | 
|
| 69597 | 118  | 
ML_val \<open>nonconst \<^term>\<open>\<forall>P (a::'a). P a\<close>\<close>  | 
119  | 
ML_val \<open>nonconst \<^term>\<open>THE x::'a. P x\<close>\<close>  | 
|
120  | 
ML_val \<open>nonconst \<^term>\<open>SOME x::'a. P x\<close>\<close>  | 
|
121  | 
ML_val \<open>nonconst \<^term>\<open>(\<lambda>A B x::'a. A x \<or> B x) = myunion\<close>\<close>  | 
|
122  | 
ML_val \<open>nonconst \<^term>\<open>(\<lambda>x::'a. False) = (\<lambda>x::'a. True)\<close>\<close>  | 
|
123  | 
ML_val \<open>nonconst \<^prop>\<open>\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> a \<in> f \<and> a \<in> g \<longrightarrow> F h\<close>\<close>  | 
|
| 33197 | 124  | 
|
| 69597 | 125  | 
ML_val \<open>mono \<^prop>\<open>Q (\<forall>x::'a set. P x)\<close>\<close>  | 
126  | 
ML_val \<open>mono \<^prop>\<open>P (a::'a)\<close>\<close>  | 
|
127  | 
ML_val \<open>mono \<^prop>\<open>{a} = {b::'a}\<close>\<close>
 | 
|
128  | 
ML_val \<open>mono \<^prop>\<open>(\<lambda>x. x = a) = (\<lambda>y. y = (b::'a))\<close>\<close>  | 
|
129  | 
ML_val \<open>mono \<^prop>\<open>(a::'a) \<in> P \<and> P \<union> P = P\<close>\<close>  | 
|
130  | 
ML_val \<open>mono \<^prop>\<open>\<forall>F::'a set set. P\<close>\<close>  | 
|
131  | 
ML_val \<open>mono \<^prop>\<open>\<not> (\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> a \<in> f \<and> a \<in> g \<longrightarrow> F h)\<close>\<close>  | 
|
132  | 
ML_val \<open>mono \<^prop>\<open>\<not> Q (\<forall>x::'a set. P x)\<close>\<close>  | 
|
133  | 
ML_val \<open>mono \<^prop>\<open>\<not> (\<forall>x::'a. P x)\<close>\<close>  | 
|
134  | 
ML_val \<open>mono \<^prop>\<open>myall P = (P = (\<lambda>x::'a. True))\<close>\<close>  | 
|
135  | 
ML_val \<open>mono \<^prop>\<open>myall P = (P = (\<lambda>x::'a. False))\<close>\<close>  | 
|
136  | 
ML_val \<open>mono \<^prop>\<open>\<forall>x::'a. P x\<close>\<close>  | 
|
137  | 
ML_val \<open>mono \<^term>\<open>(\<lambda>A B x::'a. A x \<or> B x) \<noteq> myunion\<close>\<close>  | 
|
| 33197 | 138  | 
|
| 69597 | 139  | 
ML_val \<open>nonmono \<^prop>\<open>A = (\<lambda>x::'a. True) \<and> A = (\<lambda>x. False)\<close>\<close>  | 
140  | 
ML_val \<open>nonmono \<^prop>\<open>\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> a \<in> f \<and> a \<in> g \<longrightarrow> F h\<close>\<close>  | 
|
| 33197 | 141  | 
|
| 63167 | 142  | 
ML \<open>  | 
| 54833 | 143  | 
val preproc_timeout = seconds 5.0  | 
144  | 
val mono_timeout = seconds 1.0  | 
|
| 
41008
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
145  | 
|
| 
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
146  | 
fun is_forbidden_theorem name =  | 
| 
46711
 
f745bcc4a1e5
more explicit Long_Name operations (NB: analyzing qualifiers is inherently fragile);
 
wenzelm 
parents: 
46082 
diff
changeset
 | 
147  | 
length (Long_Name.explode name) <> 2 orelse  | 
| 
 
f745bcc4a1e5
more explicit Long_Name operations (NB: analyzing qualifiers is inherently fragile);
 
wenzelm 
parents: 
46082 
diff
changeset
 | 
148  | 
String.isPrefix "type_definition" (List.last (Long_Name.explode name)) orelse  | 
| 
 
f745bcc4a1e5
more explicit Long_Name operations (NB: analyzing qualifiers is inherently fragile);
 
wenzelm 
parents: 
46082 
diff
changeset
 | 
149  | 
String.isPrefix "arity_" (List.last (Long_Name.explode name)) orelse  | 
| 
41008
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
150  | 
String.isSuffix "_def" name orelse  | 
| 
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
151  | 
String.isSuffix "_raw" name  | 
| 
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
152  | 
|
| 
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
153  | 
fun theorems_of thy =  | 
| 
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
154  | 
filter (fn (name, th) =>  | 
| 
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
155  | 
not (is_forbidden_theorem name) andalso  | 
| 65458 | 156  | 
Thm.theory_name th = Context.theory_name thy)  | 
| 56374 | 157  | 
(Global_Theory.all_thms_of thy true)  | 
| 
41008
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
158  | 
|
| 41010 | 159  | 
fun check_formulas tsp =  | 
| 
41008
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
160  | 
let  | 
| 
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
161  | 
fun is_type_actually_monotonic T =  | 
| 41010 | 162  | 
Nitpick_Mono.formulas_monotonic hol_ctxt binarize T tsp  | 
| 
67405
 
e9ab4ad7bd15
uniform use of Standard ML op-infix -- eliminated warnings;
 
wenzelm 
parents: 
67399 
diff
changeset
 | 
163  | 
val free_Ts = fold Term.add_tfrees ((op @) tsp) [] |> map TFree  | 
| 
41008
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
164  | 
val (mono_free_Ts, nonmono_free_Ts) =  | 
| 62519 | 165  | 
Timeout.apply mono_timeout  | 
| 54833 | 166  | 
(List.partition is_type_actually_monotonic) free_Ts  | 
| 
41008
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
167  | 
in  | 
| 
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
168  | 
if not (null mono_free_Ts) then "MONO"  | 
| 
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
169  | 
else if not (null nonmono_free_Ts) then "NONMONO"  | 
| 
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
170  | 
else "NIX"  | 
| 
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
171  | 
end  | 
| 62519 | 172  | 
handle Timeout.TIMEOUT _ => "TIMEOUT"  | 
| 
41008
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
173  | 
| NOT_SUPPORTED _ => "UNSUP"  | 
| 62505 | 174  | 
| exn => if Exn.is_interrupt exn then Exn.reraise exn else "UNKNOWN"  | 
| 
41008
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
175  | 
|
| 
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
176  | 
fun check_theory thy =  | 
| 
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
177  | 
let  | 
| 
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
178  | 
val path = File.tmp_path (Context.theory_name thy ^ ".out" |> Path.explode)  | 
| 
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
179  | 
val _ = File.write path ""  | 
| 
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
180  | 
fun check_theorem (name, th) =  | 
| 
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
181  | 
let  | 
| 59582 | 182  | 
val t = th |> Thm.prop_of |> Type.legacy_freeze |> close_form  | 
| 69597 | 183  | 
val neg_t = Logic.mk_implies (t, \<^prop>\<open>False\<close>)  | 
| 41805 | 184  | 
val (nondef_ts, def_ts, _, _, _, _) =  | 
| 62519 | 185  | 
Timeout.apply preproc_timeout (preprocess_formulas hol_ctxt [])  | 
| 54833 | 186  | 
neg_t  | 
| 41010 | 187  | 
val res = name ^ ": " ^ check_formulas (nondef_ts, def_ts)  | 
| 
41008
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
188  | 
in File.append path (res ^ "\n"); writeln res end  | 
| 62519 | 189  | 
handle Timeout.TIMEOUT _ => ()  | 
| 
41008
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
190  | 
in thy |> theorems_of |> List.app check_theorem end  | 
| 63167 | 191  | 
\<close>  | 
| 
41008
 
298226ba5606
added ML code for testing entire theories for monotonicity
 
blanchet 
parents: 
41006 
diff
changeset
 | 
192  | 
|
| 41015 | 193  | 
(*  | 
| 51272 | 194  | 
ML_val {* check_theory @{theory AVL2} *}
 | 
195  | 
ML_val {* check_theory @{theory Fun} *}
 | 
|
196  | 
ML_val {* check_theory @{theory Huffman} *}
 | 
|
197  | 
ML_val {* check_theory @{theory List} *}
 | 
|
198  | 
ML_val {* check_theory @{theory Map} *}
 | 
|
199  | 
ML_val {* check_theory @{theory Relation} *}
 | 
|
| 41015 | 200  | 
*)  | 
201  | 
||
| 33197 | 202  | 
end  |