src/HOL/Nitpick_Examples/Mono_Nits.thy
author blanchet
Sun, 25 Apr 2010 00:10:30 +0200
changeset 36389 8228b3a4a2ba
parent 35812 394fe2b064cd
child 37268 8ad1e3768b4f
permissions -rw-r--r--
remove "skolemize" option from Nitpick, since Skolemization is always useful
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/Nitpick_Examples/Mono_Nits.thy
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
35076
cc19e2aef17e added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents: 35071
diff changeset
     3
    Copyright   2009, 2010
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
     4
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
     5
Examples featuring Nitpick's monotonicity check.
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
     6
*)
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
     7
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
     8
header {* Examples Featuring Nitpick's Monotonicity Check *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
     9
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    10
theory Mono_Nits
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    11
imports Main
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    12
begin
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    13
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    14
ML {*
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    15
exception FAIL
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    16
35339
34819133c75e make example compile
blanchet
parents: 35191
diff changeset
    17
val subst = []
34819133c75e make example compile
blanchet
parents: 35191
diff changeset
    18
val defs = Nitpick_HOL.all_axioms_of @{theory} subst |> #1
34819133c75e make example compile
blanchet
parents: 35191
diff changeset
    19
val def_table = Nitpick_HOL.const_def_table @{context} subst defs
35071
3df45b0ce819 adapted example following previous Nitpick change and fixed minor optimization in Nitpick
blanchet
parents: 34982
diff changeset
    20
val hol_ctxt : Nitpick_HOL.hol_context =
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    21
  {thy = @{theory}, ctxt = @{context}, max_bisim_depth = ~1, boxes = [],
34982
7b8c366e34a2 added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents: 34124
diff changeset
    22
   stds = [(NONE, true)], wfs = [], user_axioms = NONE, debug = false,
7b8c366e34a2 added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents: 34124
diff changeset
    23
   binary_ints = SOME false, destroy_constrs = false, specialize = false,
36389
8228b3a4a2ba remove "skolemize" option from Nitpick, since Skolemization is always useful
blanchet
parents: 35812
diff changeset
    24
   star_linear_preds = false, fast_descrs = false, tac_timeout = NONE,
8228b3a4a2ba remove "skolemize" option from Nitpick, since Skolemization is always useful
blanchet
parents: 35812
diff changeset
    25
   evals = [], case_names = [], def_table = def_table,
8228b3a4a2ba remove "skolemize" option from Nitpick, since Skolemization is always useful
blanchet
parents: 35812
diff changeset
    26
   nondef_table = Symtab.empty, user_nondefs = [],
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33698
diff changeset
    27
   simp_table = Unsynchronized.ref Symtab.empty, psimp_table = Symtab.empty,
35807
e4d1b5cbd429 added support for "specification" and "ax_specification" constructs to Nitpick
blanchet
parents: 35386
diff changeset
    28
   choice_spec_table = Symtab.empty, intro_table = Symtab.empty,
e4d1b5cbd429 added support for "specification" and "ax_specification" constructs to Nitpick
blanchet
parents: 35386
diff changeset
    29
   ground_thm_table = Inttab.empty, ersatz_table = [],
e4d1b5cbd429 added support for "specification" and "ax_specification" constructs to Nitpick
blanchet
parents: 35386
diff changeset
    30
   skolems = Unsynchronized.ref [], special_funs = Unsynchronized.ref [],
e4d1b5cbd429 added support for "specification" and "ax_specification" constructs to Nitpick
blanchet
parents: 35386
diff changeset
    31
   unrolled_preds = Unsynchronized.ref [], wf_cache = Unsynchronized.ref [],
e4d1b5cbd429 added support for "specification" and "ax_specification" constructs to Nitpick
blanchet
parents: 35386
diff changeset
    32
   constr_cache = Unsynchronized.ref []}
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    33
(* term -> bool *)
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35807
diff changeset
    34
fun is_mono t =
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35807
diff changeset
    35
  Nitpick_Mono.formulas_monotonic hol_ctxt false @{typ 'a} ([t], [])
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    36
fun is_const t =
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    37
  let val T = fastype_of t in
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    38
    is_mono (Logic.mk_implies (Logic.mk_equals (Free ("dummyP", T), t),
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    39
                               @{const False}))
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    40
  end
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    41
fun mono t = is_mono t orelse raise FAIL
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    42
fun nonmono t = not (is_mono t) orelse raise FAIL
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    43
fun const t = is_const t orelse raise FAIL
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    44
fun nonconst t = not (is_const t) orelse raise FAIL
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    45
*}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    46
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    47
ML {* const @{term "A::('a\<Rightarrow>'b)"} *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    48
ML {* const @{term "(A::'a set) = A"} *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    49
ML {* const @{term "(A::'a set set) = A"} *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    50
ML {* const @{term "(\<lambda>x::'a set. x a)"} *}
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
    51
ML {* const @{term "{{a::'a}} = C"} *}
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    52
ML {* const @{term "{f::'a\<Rightarrow>nat} = {g::'a\<Rightarrow>nat}"} *}
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
    53
ML {* const @{term "A \<union> (B::'a set)"} *}
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    54
ML {* const @{term "P (a::'a)"} *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    55
ML {* const @{term "\<lambda>a::'a. b (c (d::'a)) (e::'a) (f::'a)"} *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    56
ML {* const @{term "\<forall>A::'a set. A a"} *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    57
ML {* const @{term "\<forall>A::'a set. P A"} *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    58
ML {* const @{term "P \<or> Q"} *}
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
    59
ML {* const @{term "A \<union> B = (C::'a set)"} *}
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    60
ML {* const @{term "(if P then (A::'a set) else B) = C"} *}
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
    61
ML {* const @{term "let A = (C::'a set) in A \<union> B"} *}
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    62
ML {* const @{term "THE x::'b. P x"} *}
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
    63
ML {* const @{term "(\<lambda>x::'a. False)"} *}
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    64
ML {* const @{term "(\<lambda>x::'a. True)"} *}
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
    65
ML {* const @{term "Let (a::'a) A"} *}
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    66
ML {* const @{term "A (a::'a)"} *}
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
    67
ML {* const @{term "insert (a::'a) A = B"} *}
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    68
ML {* const @{term "- (A::'a set)"} *}
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
    69
ML {* const @{term "finite (A::'a set)"} *}
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
    70
ML {* const @{term "\<not> finite (A::'a set)"} *}
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    71
ML {* const @{term "finite (A::'a set set)"} *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    72
ML {* const @{term "\<lambda>a::'a. A a \<and> \<not> B a"} *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    73
ML {* const @{term "A < (B::'a set)"} *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    74
ML {* const @{term "A \<le> (B::'a set)"} *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    75
ML {* const @{term "[a::'a]"} *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    76
ML {* const @{term "[a::'a set]"} *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    77
ML {* const @{term "[A \<union> (B::'a set)]"} *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    78
ML {* const @{term "[A \<union> (B::'a set)] = [C]"} *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    79
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
    80
ML {* nonconst @{term "{(\<lambda>x::'a. x = a)} = C"} *}
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    81
ML {* nonconst @{term "\<forall>P (a::'a). P a"} *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    82
ML {* nonconst @{term "\<forall>a::'a. P a"} *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    83
ML {* nonconst @{term "(\<lambda>a::'a. \<not> A a) = B"} *}
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
    84
ML {* nonconst @{term "THE x::'a. P x"} *}
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
    85
ML {* nonconst @{term "SOME x::'a. P x"} *}
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    86
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    87
ML {* mono @{prop "Q (\<forall>x::'a set. P x)"} *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    88
ML {* mono @{prop "P (a::'a)"} *}
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
    89
ML {* mono @{prop "{a} = {b::'a}"} *}
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    90
ML {* mono @{prop "P (a::'a) \<and> P \<union> P = P"} *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    91
ML {* mono @{prop "\<forall>F::'a set set. P"} *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    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)"} *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    93
ML {* mono @{prop "\<not> Q (\<forall>x::'a set. P x)"} *}
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
    94
ML {* mono @{prop "\<not> (\<forall>x::'a. P x)"} *}
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    95
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
    96
ML {* nonmono @{prop "\<forall>x::'a. P x"} *}
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
    97
ML {* nonmono @{prop "myall P = (P = (\<lambda>x::'a. True))"} *}
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    98
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"} *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    99
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   100
end