28 intro_table = Symtab.empty, ground_thm_table = Inttab.empty, |
28 intro_table = Symtab.empty, ground_thm_table = Inttab.empty, |
29 ersatz_table = [], skolems = Unsynchronized.ref [], |
29 ersatz_table = [], skolems = Unsynchronized.ref [], |
30 special_funs = Unsynchronized.ref [], unrolled_preds = Unsynchronized.ref [], |
30 special_funs = Unsynchronized.ref [], unrolled_preds = Unsynchronized.ref [], |
31 wf_cache = Unsynchronized.ref [], constr_cache = Unsynchronized.ref []} |
31 wf_cache = Unsynchronized.ref [], constr_cache = Unsynchronized.ref []} |
32 (* term -> bool *) |
32 (* term -> bool *) |
33 val is_mono = Nitpick_Mono.formulas_monotonic hol_ctxt false @{typ 'a} [] [] |
33 fun is_mono t = Nitpick_Mono.formulas_monotonic hol_ctxt false @{typ 'a} ([t], []) |
34 fun is_const t = |
34 fun is_const t = |
35 let val T = fastype_of t in |
35 let val T = fastype_of t in |
36 is_mono (Logic.mk_implies (Logic.mk_equals (Free ("dummyP", T), t), |
36 is_mono (Logic.mk_implies (Logic.mk_equals (Free ("dummyP", T), t), |
37 @{const False})) |
37 @{const False})) |
38 end |
38 end |
44 |
44 |
45 ML {* const @{term "A::('a\<Rightarrow>'b)"} *} |
45 ML {* const @{term "A::('a\<Rightarrow>'b)"} *} |
46 ML {* const @{term "(A::'a set) = A"} *} |
46 ML {* const @{term "(A::'a set) = A"} *} |
47 ML {* const @{term "(A::'a set set) = A"} *} |
47 ML {* const @{term "(A::'a set set) = A"} *} |
48 ML {* const @{term "(\<lambda>x::'a set. x a)"} *} |
48 ML {* const @{term "(\<lambda>x::'a set. x a)"} *} |
49 ML {* const @{term "{{a}} = C"} *} |
49 ML {* const @{term "{{a::'a}} = C"} *} |
50 ML {* const @{term "{f::'a\<Rightarrow>nat} = {g::'a\<Rightarrow>nat}"} *} |
50 ML {* const @{term "{f::'a\<Rightarrow>nat} = {g::'a\<Rightarrow>nat}"} *} |
51 ML {* const @{term "A \<union> B"} *} |
51 ML {* const @{term "A \<union> (B::'a set)"} *} |
52 ML {* const @{term "P (a::'a)"} *} |
52 ML {* const @{term "P (a::'a)"} *} |
53 ML {* const @{term "\<lambda>a::'a. b (c (d::'a)) (e::'a) (f::'a)"} *} |
53 ML {* const @{term "\<lambda>a::'a. b (c (d::'a)) (e::'a) (f::'a)"} *} |
54 ML {* const @{term "\<forall>A::'a set. A a"} *} |
54 ML {* const @{term "\<forall>A::'a set. A a"} *} |
55 ML {* const @{term "\<forall>A::'a set. P A"} *} |
55 ML {* const @{term "\<forall>A::'a set. P A"} *} |
56 ML {* const @{term "P \<or> Q"} *} |
56 ML {* const @{term "P \<or> Q"} *} |
57 ML {* const @{term "A \<union> B = C"} *} |
57 ML {* const @{term "A \<union> B = (C::'a set)"} *} |
58 ML {* const @{term "(if P then (A::'a set) else B) = C"} *} |
58 ML {* const @{term "(if P then (A::'a set) else B) = C"} *} |
59 ML {* const @{term "let A = C in A \<union> B"} *} |
59 ML {* const @{term "let A = (C::'a set) in A \<union> B"} *} |
60 ML {* const @{term "THE x::'b. P x"} *} |
60 ML {* const @{term "THE x::'b. P x"} *} |
61 ML {* const @{term "{}::'a set"} *} |
61 ML {* const @{term "(\<lambda>x::'a. False)"} *} |
62 ML {* const @{term "(\<lambda>x::'a. True)"} *} |
62 ML {* const @{term "(\<lambda>x::'a. True)"} *} |
63 ML {* const @{term "Let a A"} *} |
63 ML {* const @{term "Let (a::'a) A"} *} |
64 ML {* const @{term "A (a::'a)"} *} |
64 ML {* const @{term "A (a::'a)"} *} |
65 ML {* const @{term "insert a A = B"} *} |
65 ML {* const @{term "insert (a::'a) A = B"} *} |
66 ML {* const @{term "- (A::'a set)"} *} |
66 ML {* const @{term "- (A::'a set)"} *} |
67 ML {* const @{term "finite A"} *} |
67 ML {* const @{term "finite (A::'a set)"} *} |
68 ML {* const @{term "\<not> finite A"} *} |
68 ML {* const @{term "\<not> finite (A::'a set)"} *} |
69 ML {* const @{term "finite (A::'a set set)"} *} |
69 ML {* const @{term "finite (A::'a set set)"} *} |
70 ML {* const @{term "\<lambda>a::'a. A a \<and> \<not> B a"} *} |
70 ML {* const @{term "\<lambda>a::'a. A a \<and> \<not> B a"} *} |
71 ML {* const @{term "A < (B::'a set)"} *} |
71 ML {* const @{term "A < (B::'a set)"} *} |
72 ML {* const @{term "A \<le> (B::'a set)"} *} |
72 ML {* const @{term "A \<le> (B::'a set)"} *} |
73 ML {* const @{term "[a::'a]"} *} |
73 ML {* const @{term "[a::'a]"} *} |
74 ML {* const @{term "[a::'a set]"} *} |
74 ML {* const @{term "[a::'a set]"} *} |
75 ML {* const @{term "[A \<union> (B::'a set)]"} *} |
75 ML {* const @{term "[A \<union> (B::'a set)]"} *} |
76 ML {* const @{term "[A \<union> (B::'a set)] = [C]"} *} |
76 ML {* const @{term "[A \<union> (B::'a set)] = [C]"} *} |
77 ML {* const @{term "\<forall>P. P a"} *} |
|
78 |
77 |
79 ML {* nonconst @{term "{%x. True}"} *} |
78 ML {* nonconst @{term "{(\<lambda>x::'a. x = a)} = C"} *} |
80 ML {* nonconst @{term "{(%x. x = a)} = C"} *} |
|
81 ML {* nonconst @{term "\<forall>P (a::'a). P a"} *} |
79 ML {* nonconst @{term "\<forall>P (a::'a). P a"} *} |
82 ML {* nonconst @{term "\<forall>a::'a. P a"} *} |
80 ML {* nonconst @{term "\<forall>a::'a. P a"} *} |
83 ML {* nonconst @{term "(\<lambda>a::'a. \<not> A a) = B"} *} |
81 ML {* nonconst @{term "(\<lambda>a::'a. \<not> A a) = B"} *} |
84 ML {* nonconst @{term "THE x. P x"} *} |
82 ML {* nonconst @{term "THE x::'a. P x"} *} |
85 ML {* nonconst @{term "SOME x. P x"} *} |
83 ML {* nonconst @{term "SOME x::'a. P x"} *} |
86 |
84 |
87 ML {* mono @{prop "Q (\<forall>x::'a set. P x)"} *} |
85 ML {* mono @{prop "Q (\<forall>x::'a set. P x)"} *} |
88 ML {* mono @{prop "P (a::'a)"} *} |
86 ML {* mono @{prop "P (a::'a)"} *} |
89 ML {* mono @{prop "{a} = {b}"} *} |
87 ML {* mono @{prop "{a} = {b::'a}"} *} |
90 ML {* mono @{prop "P (a::'a) \<and> P \<union> P = P"} *} |
88 ML {* mono @{prop "P (a::'a) \<and> P \<union> P = P"} *} |
91 ML {* mono @{prop "\<forall>F::'a set set. P"} *} |
89 ML {* mono @{prop "\<forall>F::'a set set. P"} *} |
92 ML {* mono @{prop "\<not> (\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> f a \<and> g a \<longrightarrow> F h)"} *} |
90 ML {* mono @{prop "\<not> (\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> f a \<and> g a \<longrightarrow> F h)"} *} |
93 ML {* mono @{prop "\<not> Q (\<forall>x::'a set. P x)"} *} |
91 ML {* mono @{prop "\<not> Q (\<forall>x::'a set. P x)"} *} |
94 ML {* mono @{prop "\<not> (\<forall>x. P x)"} *} |
92 ML {* mono @{prop "\<not> (\<forall>x::'a. P x)"} *} |
95 |
93 |
96 ML {* nonmono @{prop "\<forall>x. P x"} *} |
94 ML {* nonmono @{prop "\<forall>x::'a. P x"} *} |
|
95 ML {* nonmono @{prop "myall P = (P = (\<lambda>x::'a. True))"} *} |
97 ML {* nonmono @{prop "\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> f a \<and> g a \<longrightarrow> F h"} *} |
96 ML {* nonmono @{prop "\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> f a \<and> g a \<longrightarrow> F h"} *} |
98 ML {* nonmono @{prop "myall P = (P = (\<lambda>x. True))"} *} |
|
99 |
97 |
100 end |
98 end |