src/HOL/Tools/Nitpick/nitpick_model.ML
author blanchet
Tue, 14 Sep 2010 13:24:18 +0200
changeset 39359 6f49c7fbb1b1
parent 39345 062c10ff848c
child 39360 cdf2c3341422
permissions -rw-r--r--
remove "fast_descs" option from Nitpick; the option has been unsound for over a year and is too imprecise to be of any use when made sound
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33982
1ae222745c4a fixed paths in Nitpick's ML file headers
blanchet
parents: 33884
diff changeset
     1
(*  Title:      HOL/Tools/Nitpick/nitpick_model.ML
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
34982
7b8c366e34a2 added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents: 34974
diff changeset
     3
    Copyright   2009, 2010
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
     4
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
     5
Model reconstruction for Nitpick.
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
     6
*)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
     7
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
     8
signature NITPICK_MODEL =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
     9
sig
33705
947184dc75c9 removed a few global names in Nitpick (styp, nat_less, pairf)
blanchet
parents: 33580
diff changeset
    10
  type styp = Nitpick_Util.styp
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33202
diff changeset
    11
  type scope = Nitpick_Scope.scope
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33202
diff changeset
    12
  type rep = Nitpick_Rep.rep
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33202
diff changeset
    13
  type nut = Nitpick_Nut.nut
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    14
36390
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36389
diff changeset
    15
  type params =
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36389
diff changeset
    16
    {show_datatypes: bool,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36389
diff changeset
    17
     show_consts: bool}
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36389
diff changeset
    18
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
    19
  type term_postprocessor =
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
    20
    Proof.context -> string -> (typ -> term list) -> typ -> term -> term
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    21
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    22
  structure NameTable : TABLE
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    23
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
    24
  val irrelevant : string
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
    25
  val unknown : string
37261
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
    26
  val unrep : unit -> string
35711
548d3f16404b added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents: 35665
diff changeset
    27
  val register_term_postprocessor :
38284
9f98107ad8b4 use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents: 38240
diff changeset
    28
    typ -> term_postprocessor -> morphism -> Context.generic -> Context.generic
38240
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38209
diff changeset
    29
  val register_term_postprocessor_global :
35711
548d3f16404b added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents: 35665
diff changeset
    30
    typ -> term_postprocessor -> theory -> theory
38284
9f98107ad8b4 use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents: 38240
diff changeset
    31
  val unregister_term_postprocessor :
9f98107ad8b4 use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents: 38240
diff changeset
    32
    typ -> morphism -> Context.generic -> Context.generic
38240
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38209
diff changeset
    33
  val unregister_term_postprocessor_global : typ -> theory -> theory
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    34
  val tuple_list_for_name :
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    35
    nut NameTable.table -> Kodkod.raw_bound list -> nut -> int list list
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
    36
  val dest_plain_fun : term -> bool * (term list * term list)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    37
  val reconstruct_hol_model :
37260
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
    38
    params -> scope -> (term option * int list) list
38170
d74b66ec02ce handle free variables even more gracefully;
blanchet
parents: 38126
diff changeset
    39
    -> (typ option * string list) list -> styp list -> styp list -> nut list
d74b66ec02ce handle free variables even more gracefully;
blanchet
parents: 38126
diff changeset
    40
    -> nut list -> nut list -> nut NameTable.table -> Kodkod.raw_bound list
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    41
    -> Pretty.T * bool
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    42
  val prove_hol_model :
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    43
    scope -> Time.time option -> nut list -> nut list -> nut NameTable.table
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    44
    -> Kodkod.raw_bound list -> term -> bool option
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    45
end;
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    46
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33202
diff changeset
    47
structure Nitpick_Model : NITPICK_MODEL =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    48
struct
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    49
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33202
diff changeset
    50
open Nitpick_Util
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33202
diff changeset
    51
open Nitpick_HOL
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33202
diff changeset
    52
open Nitpick_Scope
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33202
diff changeset
    53
open Nitpick_Peephole
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33202
diff changeset
    54
open Nitpick_Rep
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33202
diff changeset
    55
open Nitpick_Nut
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    56
34126
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 34124
diff changeset
    57
structure KK = Kodkod
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 34124
diff changeset
    58
36390
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36389
diff changeset
    59
type params =
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36389
diff changeset
    60
  {show_datatypes: bool,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36389
diff changeset
    61
   show_consts: bool}
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    62
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
    63
type term_postprocessor =
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
    64
  Proof.context -> string -> (typ -> term list) -> typ -> term -> term
35711
548d3f16404b added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents: 35665
diff changeset
    65
38240
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38209
diff changeset
    66
structure Data = Generic_Data(
35711
548d3f16404b added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents: 35665
diff changeset
    67
  type T = (typ * term_postprocessor) list
548d3f16404b added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents: 35665
diff changeset
    68
  val empty = []
548d3f16404b added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents: 35665
diff changeset
    69
  val extend = I
36607
e5f7235f39c5 made sml/nj happy about Sledgehammer and Nitpick (cf. 6f11c9b1fb3e, 3c2438efe224)
krauss
parents: 36391
diff changeset
    70
  fun merge (x, y) = AList.merge (op =) (K true) (x, y))
35711
548d3f16404b added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents: 35665
diff changeset
    71
37261
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
    72
fun xsym s s' () = if print_mode_active Symbol.xsymbolsN then s else s'
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
    73
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
    74
val irrelevant = "_"
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    75
val unknown = "?"
37261
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
    76
val unrep = xsym "\<dots>" "..."
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
    77
val maybe_mixfix = xsym "_\<^sup>?" "_?"
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
    78
val base_mixfix = xsym "_\<^bsub>base\<^esub>" "_.base"
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
    79
val step_mixfix = xsym "_\<^bsub>step\<^esub>" "_.step"
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
    80
val abs_mixfix = xsym "\<guillemotleft>_\<guillemotright>" "\"_\""
35718
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
    81
val arg_var_prefix = "x"
37261
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
    82
val cyclic_co_val_name = xsym "\<omega>" "w"
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
    83
val cyclic_const_prefix = xsym "\<xi>" "X"
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
    84
fun cyclic_type_name () = nitpick_prefix ^ cyclic_const_prefix ()
34982
7b8c366e34a2 added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents: 34974
diff changeset
    85
val opt_flag = nitpick_prefix ^ "opt"
7b8c366e34a2 added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents: 34974
diff changeset
    86
val non_opt_flag = nitpick_prefix ^ "non_opt"
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    87
35076
cc19e2aef17e added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents: 35075
diff changeset
    88
type atom_pool = ((string * int) * int list) list
cc19e2aef17e added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents: 35075
diff changeset
    89
35718
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
    90
fun add_wacky_syntax ctxt =
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
    91
  let
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
    92
    val name_of = fst o dest_Const
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
    93
    val thy = ProofContext.theory_of ctxt |> Context.reject_draft
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
    94
    val (maybe_t, thy) =
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
    95
      Sign.declare_const ((@{binding nitpick_maybe}, @{typ "'a => 'a"}),
37261
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
    96
                          Mixfix (maybe_mixfix (), [1000], 1000)) thy
35718
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
    97
    val (abs_t, thy) =
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
    98
      Sign.declare_const ((@{binding nitpick_abs}, @{typ "'a => 'b"}),
37261
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
    99
                          Mixfix (abs_mixfix (), [40], 40)) thy
35718
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   100
    val (base_t, thy) =
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   101
      Sign.declare_const ((@{binding nitpick_base}, @{typ "'a => 'a"}),
37261
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
   102
                          Mixfix (base_mixfix (), [1000], 1000)) thy
35718
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   103
    val (step_t, thy) =
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   104
      Sign.declare_const ((@{binding nitpick_step}, @{typ "'a => 'a"}),
37261
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
   105
                          Mixfix (step_mixfix (), [1000], 1000)) thy
35718
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   106
  in
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   107
    (pairself (pairself name_of) ((maybe_t, abs_t), (base_t, step_t)),
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   108
     ProofContext.transfer_syntax thy ctxt)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   109
  end
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   110
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   111
(** Term reconstruction **)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   112
37260
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   113
fun nth_atom_number pool T j =
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   114
  case AList.lookup (op =) (!pool) T of
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   115
    SOME js =>
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   116
    (case find_index (curry (op =) j) js of
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   117
       ~1 => (Unsynchronized.change pool (cons (T, j :: js));
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   118
              length js + 1)
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   119
     | n => length js - n)
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   120
  | NONE => (Unsynchronized.change pool (cons (T, [j])); 1)
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   121
fun atom_suffix s =
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   122
  nat_subscript
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   123
  #> (s <> "" andalso Symbol.is_ascii_digit (List.last (explode s)))
33884
a0c43f185fef generate clearer atom names in Nitpick for types that end with a digit;
blanchet
parents: 33705
diff changeset
   124
     ? prefix "\<^isub>,"
37260
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   125
fun nth_atom_name thy atomss pool prefix T j =
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   126
  let
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   127
    val ss = these (triple_lookup (type_match thy) atomss T)
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   128
    val m = nth_atom_number pool T j
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   129
  in
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   130
    if m <= length ss then
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   131
      nth ss (m - 1)
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   132
    else case T of
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   133
      Type (s, _) =>
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   134
      let val s' = shortest_name s in
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   135
        prefix ^
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   136
        (if String.isPrefix "\\" s' then s' else substring (s', 0, 1)) ^
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   137
        atom_suffix s m
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   138
      end
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   139
    | TFree (s, _) => prefix ^ perhaps (try (unprefix "'")) s ^ atom_suffix s m
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   140
    | _ => raise TYPE ("Nitpick_Model.nth_atom_name", [T], [])
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   141
  end
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   142
fun nth_atom thy atomss pool for_auto T j =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   143
  if for_auto then
37260
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   144
    Free (nth_atom_name thy atomss pool (hd (space_explode "." nitpick_prefix))
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   145
                        T j, T)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   146
  else
37260
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   147
    Const (nth_atom_name thy atomss pool "" T j, T)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   148
35177
168041f24f80 various cosmetic changes to Nitpick
blanchet
parents: 35086
diff changeset
   149
fun extract_real_number (Const (@{const_name divide}, _) $ t1 $ t2) =
34126
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 34124
diff changeset
   150
    real (snd (HOLogic.dest_number t1)) / real (snd (HOLogic.dest_number t2))
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 34124
diff changeset
   151
  | extract_real_number t = real (snd (HOLogic.dest_number t))
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   152
fun nice_term_ord (Abs (_, _, t1), Abs (_, _, t2)) = nice_term_ord (t1, t2)
34126
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 34124
diff changeset
   153
  | nice_term_ord tp = Real.compare (pairself extract_real_number tp)
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   154
    handle TERM ("dest_number", _) =>
34126
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 34124
diff changeset
   155
           case tp of
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   156
             (t11 $ t12, t21 $ t22) =>
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   157
             (case nice_term_ord (t11, t21) of
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   158
                EQUAL => nice_term_ord (t12, t22)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   159
              | ord => ord)
35408
b48ab741683b modernized structure Term_Ord;
wenzelm
parents: 35402
diff changeset
   160
           | _ => Term_Ord.fast_term_ord tp
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   161
38284
9f98107ad8b4 use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents: 38240
diff changeset
   162
fun register_term_postprocessor_generic T postproc =
9f98107ad8b4 use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents: 38240
diff changeset
   163
  Data.map (cons (T, postproc))
9f98107ad8b4 use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents: 38240
diff changeset
   164
(* TODO: Consider morphism. *)
9f98107ad8b4 use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents: 38240
diff changeset
   165
fun register_term_postprocessor T postproc (_ : morphism) =
9f98107ad8b4 use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents: 38240
diff changeset
   166
  register_term_postprocessor_generic T postproc
38240
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38209
diff changeset
   167
val register_term_postprocessor_global =
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38209
diff changeset
   168
  Context.theory_map oo register_term_postprocessor_generic
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38209
diff changeset
   169
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38209
diff changeset
   170
fun unregister_term_postprocessor_generic T = Data.map (AList.delete (op =) T)
38284
9f98107ad8b4 use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents: 38240
diff changeset
   171
(* TODO: Consider morphism. *)
9f98107ad8b4 use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents: 38240
diff changeset
   172
fun unregister_term_postprocessor T (_ : morphism) =
9f98107ad8b4 use "declaration" instead of "setup" to register Nitpick extensions
blanchet
parents: 38240
diff changeset
   173
  unregister_term_postprocessor_generic T
38240
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38209
diff changeset
   174
val unregister_term_postprocessor_global =
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38209
diff changeset
   175
  Context.theory_map o unregister_term_postprocessor_generic
35711
548d3f16404b added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents: 35665
diff changeset
   176
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   177
fun tuple_list_for_name rel_table bounds name =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   178
  the (AList.lookup (op =) bounds (the_rel rel_table name)) handle NUT _ => [[]]
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   179
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   180
fun unarize_unbox_etc_term (Const (@{const_name FinFun}, _) $ t1) =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   181
    unarize_unbox_etc_term t1
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   182
  | unarize_unbox_etc_term (Const (@{const_name FunBox}, _) $ t1) =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   183
    unarize_unbox_etc_term t1
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   184
  | unarize_unbox_etc_term
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   185
        (Const (@{const_name PairBox},
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   186
                Type (@{type_name fun}, [T1, Type (@{type_name fun}, [T2, _])]))
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   187
         $ t1 $ t2) =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   188
    let val Ts = map uniterize_unarize_unbox_etc_type [T1, T2] in
38190
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 38189
diff changeset
   189
      Const (@{const_name Pair}, Ts ---> Type (@{type_name prod}, Ts))
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   190
      $ unarize_unbox_etc_term t1 $ unarize_unbox_etc_term t2
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   191
    end
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   192
  | unarize_unbox_etc_term (Const (s, T)) =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   193
    Const (s, uniterize_unarize_unbox_etc_type T)
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   194
  | unarize_unbox_etc_term (t1 $ t2) =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   195
    unarize_unbox_etc_term t1 $ unarize_unbox_etc_term t2
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   196
  | unarize_unbox_etc_term (Free (s, T)) =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   197
    Free (s, uniterize_unarize_unbox_etc_type T)
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   198
  | unarize_unbox_etc_term (Var (x, T)) =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   199
    Var (x, uniterize_unarize_unbox_etc_type T)
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   200
  | unarize_unbox_etc_term (Bound j) = Bound j
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   201
  | unarize_unbox_etc_term (Abs (s, T, t')) =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   202
    Abs (s, uniterize_unarize_unbox_etc_type T, unarize_unbox_etc_term t')
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   203
38190
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 38189
diff changeset
   204
fun factor_out_types (T1 as Type (@{type_name prod}, [T11, T12]))
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 38189
diff changeset
   205
                     (T2 as Type (@{type_name prod}, [T21, T22])) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   206
    let val (n1, n2) = pairself num_factors_in_type (T11, T21) in
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   207
      if n1 = n2 then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   208
        let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   209
          val ((T11', opt_T12'), (T21', opt_T22')) = factor_out_types T12 T22
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   210
        in
38190
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 38189
diff changeset
   211
          ((Type (@{type_name prod}, [T11, T11']), opt_T12'),
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 38189
diff changeset
   212
           (Type (@{type_name prod}, [T21, T21']), opt_T22'))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   213
        end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   214
      else if n1 < n2 then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   215
        case factor_out_types T1 T21 of
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   216
          (p1, (T21', NONE)) => (p1, (T21', SOME T22))
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   217
        | (p1, (T21', SOME T22')) =>
38190
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 38189
diff changeset
   218
          (p1, (T21', SOME (Type (@{type_name prod}, [T22', T22]))))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   219
      else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   220
        swap (factor_out_types T2 T1)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   221
    end
38190
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 38189
diff changeset
   222
  | factor_out_types (Type (@{type_name prod}, [T11, T12])) T2 =
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   223
    ((T11, SOME T12), (T2, NONE))
38190
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 38189
diff changeset
   224
  | factor_out_types T1 (Type (@{type_name prod}, [T21, T22])) =
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   225
    ((T1, NONE), (T21, SOME T22))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   226
  | factor_out_types T1 T2 = ((T1, NONE), (T2, NONE))
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   227
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   228
fun make_plain_fun maybe_opt T1 T2 =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   229
  let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   230
    fun aux T1 T2 [] =
34982
7b8c366e34a2 added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents: 34974
diff changeset
   231
        Const (if maybe_opt then opt_flag else non_opt_flag, T1 --> T2)
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   232
      | aux T1 T2 ((t1, t2) :: tps) =
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   233
        Const (@{const_name fun_upd}, (T1 --> T2) --> T1 --> T2 --> T1 --> T2)
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   234
        $ aux T1 T2 tps $ t1 $ t2
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   235
  in aux T1 T2 o rev end
34982
7b8c366e34a2 added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents: 34974
diff changeset
   236
fun is_plain_fun (Const (s, _)) = (s = opt_flag orelse s = non_opt_flag)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   237
  | is_plain_fun (Const (@{const_name fun_upd}, _) $ t0 $ _ $ _) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   238
    is_plain_fun t0
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   239
  | is_plain_fun _ = false
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   240
val dest_plain_fun =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   241
  let
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   242
    fun aux (Abs (_, _, Const (s, _))) = (s <> irrelevant, ([], []))
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   243
      | aux (Const (s, _)) = (s <> non_opt_flag, ([], []))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   244
      | aux (Const (@{const_name fun_upd}, _) $ t0 $ t1 $ t2) =
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   245
        let val (maybe_opt, (ts1, ts2)) = aux t0 in
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   246
          (maybe_opt, (t1 :: ts1, t2 :: ts2))
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   247
        end
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33202
diff changeset
   248
      | aux t = raise TERM ("Nitpick_Model.dest_plain_fun", [t])
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   249
  in apsnd (pairself rev) o aux end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   250
33565
5fad8e36dfb1 fixed error in Nitpick's display of uncurried constants, which resulted in an exception
blanchet
parents: 33558
diff changeset
   251
fun break_in_two T T1 T2 t =
5fad8e36dfb1 fixed error in Nitpick's display of uncurried constants, which resulted in an exception
blanchet
parents: 33558
diff changeset
   252
  let
5fad8e36dfb1 fixed error in Nitpick's display of uncurried constants, which resulted in an exception
blanchet
parents: 33558
diff changeset
   253
    val ps = HOLogic.flat_tupleT_paths T
5fad8e36dfb1 fixed error in Nitpick's display of uncurried constants, which resulted in an exception
blanchet
parents: 33558
diff changeset
   254
    val cut = length (HOLogic.strip_tupleT T1)
5fad8e36dfb1 fixed error in Nitpick's display of uncurried constants, which resulted in an exception
blanchet
parents: 33558
diff changeset
   255
    val (ps1, ps2) = pairself HOLogic.flat_tupleT_paths (T1, T2)
5fad8e36dfb1 fixed error in Nitpick's display of uncurried constants, which resulted in an exception
blanchet
parents: 33558
diff changeset
   256
    val (ts1, ts2) = t |> HOLogic.strip_ptuple ps |> chop cut
5fad8e36dfb1 fixed error in Nitpick's display of uncurried constants, which resulted in an exception
blanchet
parents: 33558
diff changeset
   257
  in (HOLogic.mk_ptuple ps1 T1 ts1, HOLogic.mk_ptuple ps2 T2 ts2) end
38190
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 38189
diff changeset
   258
fun pair_up (Type (@{type_name prod}, [T1', T2']))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   259
            (t1 as Const (@{const_name Pair},
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   260
                          Type (@{type_name fun},
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   261
                                [_, Type (@{type_name fun}, [_, T1])]))
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   262
             $ t11 $ t12) t2 =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   263
    if T1 = T1' then HOLogic.mk_prod (t1, t2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   264
    else HOLogic.mk_prod (t11, pair_up T2' t12 t2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   265
  | pair_up _ t1 t2 = HOLogic.mk_prod (t1, t2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   266
fun multi_pair_up T1 t1 (ts2, ts3) = map2 (pair o pair_up T1 t1) ts2 ts3
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   267
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   268
fun typecast_fun (Type (@{type_name fun}, [T1', T2'])) T1 T2 t =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   269
    let
33565
5fad8e36dfb1 fixed error in Nitpick's display of uncurried constants, which resulted in an exception
blanchet
parents: 33558
diff changeset
   270
      fun do_curry T1 T1a T1b T2 t =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   271
        let
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   272
          val (maybe_opt, tsp) = dest_plain_fun t
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   273
          val tps =
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   274
            tsp |>> map (break_in_two T1 T1a T1b)
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   275
                |> uncurry (map2 (fn (t1a, t1b) => fn t2 => (t1a, (t1b, t2))))
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   276
                |> AList.coalesce (op =)
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   277
                |> map (apsnd (make_plain_fun maybe_opt T1b T2))
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   278
        in make_plain_fun maybe_opt T1a (T1b --> T2) tps end
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   279
      and do_uncurry T1 T2 t =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   280
        let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   281
          val (maybe_opt, tsp) = dest_plain_fun t
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   282
          val tps =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   283
            tsp |> op ~~
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   284
                |> maps (fn (t1, t2) =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   285
                            multi_pair_up T1 t1 (snd (dest_plain_fun t2)))
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   286
        in make_plain_fun maybe_opt T1 T2 tps end
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   287
      and do_arrow T1' T2' _ _ (Const (s, _)) = Const (s, T1' --> T2')
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   288
        | do_arrow T1' T2' T1 T2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   289
                   (Const (@{const_name fun_upd}, _) $ t0 $ t1 $ t2) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   290
          Const (@{const_name fun_upd},
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   291
                 (T1' --> T2') --> T1' --> T2' --> T1' --> T2')
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   292
          $ do_arrow T1' T2' T1 T2 t0 $ do_term T1' T1 t1 $ do_term T2' T2 t2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   293
        | do_arrow _ _ _ _ t =
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33202
diff changeset
   294
          raise TERM ("Nitpick_Model.typecast_fun.do_arrow", [t])
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   295
      and do_fun T1' T2' T1 T2 t =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   296
        case factor_out_types T1' T1 of
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   297
          ((_, NONE), (_, NONE)) => t |> do_arrow T1' T2' T1 T2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   298
        | ((_, NONE), (T1a, SOME T1b)) =>
33565
5fad8e36dfb1 fixed error in Nitpick's display of uncurried constants, which resulted in an exception
blanchet
parents: 33558
diff changeset
   299
          t |> do_curry T1 T1a T1b T2 |> do_arrow T1' T2' T1a (T1b --> T2)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   300
        | ((T1a', SOME T1b'), (_, NONE)) =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   301
          t |> do_arrow T1a' (T1b' --> T2') T1 T2 |> do_uncurry T1' T2'
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33202
diff changeset
   302
        | _ => raise TYPE ("Nitpick_Model.typecast_fun.do_fun", [T1, T1'], [])
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   303
      and do_term (Type (@{type_name fun}, [T1', T2']))
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   304
                  (Type (@{type_name fun}, [T1, T2])) t =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   305
          do_fun T1' T2' T1 T2 t
38190
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 38189
diff changeset
   306
        | do_term (T' as Type (@{type_name prod}, Ts' as [T1', T2']))
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 38189
diff changeset
   307
                  (Type (@{type_name prod}, [T1, T2]))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   308
                  (Const (@{const_name Pair}, _) $ t1 $ t2) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   309
          Const (@{const_name Pair}, Ts' ---> T')
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   310
          $ do_term T1' T1 t1 $ do_term T2' T2 t2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   311
        | do_term T' T t =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   312
          if T = T' then t
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33202
diff changeset
   313
          else raise TYPE ("Nitpick_Model.typecast_fun.do_term", [T, T'], [])
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   314
    in if T1' = T1 andalso T2' = T2 then t else do_fun T1' T2' T1 T2 t end
34998
5e492a862b34 four changes to Nitpick:
blanchet
parents: 34982
diff changeset
   315
  | typecast_fun T' _ _ _ =
5e492a862b34 four changes to Nitpick:
blanchet
parents: 34982
diff changeset
   316
    raise TYPE ("Nitpick_Model.typecast_fun", [T'], [])
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   317
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   318
fun truth_const_sort_key @{const True} = "0"
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   319
  | truth_const_sort_key @{const False} = "2"
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   320
  | truth_const_sort_key _ = "1"
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   321
38190
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 38189
diff changeset
   322
fun mk_tuple (Type (@{type_name prod}, [T1, T2])) ts =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   323
    HOLogic.mk_prod (mk_tuple T1 ts,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   324
        mk_tuple T2 (List.drop (ts, length (HOLogic.flatten_tupleT T1))))
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   325
  | mk_tuple _ (t :: _) = t
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   326
  | mk_tuple T [] = raise TYPE ("Nitpick_Model.mk_tuple", [T], [])
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   327
38240
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38209
diff changeset
   328
fun varified_type_match ctxt (candid_T, pat_T) =
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38209
diff changeset
   329
  let val thy = ProofContext.theory_of ctxt in
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38209
diff changeset
   330
    strict_type_match thy (candid_T, varify_type ctxt pat_T)
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38209
diff changeset
   331
  end
35711
548d3f16404b added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents: 35665
diff changeset
   332
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   333
fun all_values_of_type pool wacky_names (scope as {card_assigns, ...} : scope)
37260
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   334
                       atomss sel_names rel_table bounds card T =
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   335
  let
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   336
    val card = if card = 0 then card_of_type card_assigns T else card
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   337
    fun nth_value_of_type n =
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   338
      let
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   339
        fun term unfold =
37262
c0fe8fa35771 don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents: 37261
diff changeset
   340
          reconstruct_term true unfold pool wacky_names scope atomss sel_names
37260
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   341
                           rel_table bounds T T (Atom (card, 0)) [[n]]
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   342
      in
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   343
        case term false of
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   344
          t as Const (s, _) =>
37261
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
   345
          if String.isPrefix (cyclic_const_prefix ()) s then
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   346
            HOLogic.mk_eq (t, term true)
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   347
          else
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   348
            t
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   349
        | t => t
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   350
      end
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   351
  in index_seq 0 card |> map nth_value_of_type |> sort nice_term_ord end
37262
c0fe8fa35771 don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents: 37261
diff changeset
   352
and reconstruct_term maybe_opt unfold pool
c0fe8fa35771 don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents: 37261
diff changeset
   353
        (wacky_names as ((maybe_name, abs_name), _))
37256
0dca1ec52999 thread along context instead of theory for typedef lookup
blanchet
parents: 37170
diff changeset
   354
        (scope as {hol_ctxt as {thy, ctxt, stds, ...}, binarize, card_assigns,
37260
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   355
                   bits, datatypes, ofs, ...})
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   356
        atomss sel_names rel_table bounds =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   357
  let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   358
    val for_auto = (maybe_name = "")
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   359
    fun value_of_bits jss =
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   360
      let
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   361
        val j0 = offset_of_type ofs @{typ unsigned_bit}
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   362
        val js = map (Integer.add (~ j0) o the_single) jss
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   363
      in
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   364
        fold (fn j => Integer.add (reasonable_power 2 j |> j = bits ? op ~))
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   365
             js 0
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   366
      end
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   367
    val all_values =
37260
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   368
      all_values_of_type pool wacky_names scope atomss sel_names rel_table
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   369
                         bounds 0
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   370
    fun postprocess_term (Type (@{type_name fun}, _)) = I
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   371
      | postprocess_term T =
38240
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38209
diff changeset
   372
        case Data.get (Context.Proof ctxt) of
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38209
diff changeset
   373
          [] => I
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38209
diff changeset
   374
        | postprocs =>
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38209
diff changeset
   375
          case AList.lookup (varified_type_match ctxt) postprocs T of
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38209
diff changeset
   376
            SOME postproc => postproc ctxt maybe_name all_values T
a44d108a8d39 local versions of Nitpick.register_xxx functions
blanchet
parents: 38209
diff changeset
   377
          | NONE => I
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   378
    fun postprocess_subterms Ts (t1 $ t2) =
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   379
        let val t = postprocess_subterms Ts t1 $ postprocess_subterms Ts t2 in
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   380
          postprocess_term (fastype_of1 (Ts, t)) t
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   381
        end
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   382
      | postprocess_subterms Ts (Abs (s, T, t')) =
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   383
        Abs (s, T, postprocess_subterms (T :: Ts) t')
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   384
      | postprocess_subterms Ts t = postprocess_term (fastype_of1 (Ts, t)) t
35388
42d39948cace use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents: 35385
diff changeset
   385
    fun make_set maybe_opt T1 T2 tps =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   386
      let
35402
115a5a95710a clarified @{const_name} vs. @{const_abbrev};
wenzelm
parents: 35388
diff changeset
   387
        val empty_const = Const (@{const_abbrev Set.empty}, T1 --> T2)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   388
        val insert_const = Const (@{const_name insert},
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   389
                                  T1 --> (T1 --> T2) --> T1 --> T2)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   390
        fun aux [] =
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   391
            if maybe_opt andalso not (is_complete_type datatypes false T1) then
37261
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
   392
              insert_const $ Const (unrep (), T1) $ empty_const
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   393
            else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   394
              empty_const
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   395
          | aux ((t1, t2) :: zs) =
35388
42d39948cace use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents: 35385
diff changeset
   396
            aux zs
42d39948cace use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents: 35385
diff changeset
   397
            |> t2 <> @{const False}
42d39948cace use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents: 35385
diff changeset
   398
               ? curry (op $)
42d39948cace use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents: 35385
diff changeset
   399
                       (insert_const
42d39948cace use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents: 35385
diff changeset
   400
                        $ (t1 |> t2 <> @{const True}
42d39948cace use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents: 35385
diff changeset
   401
                                 ? curry (op $)
42d39948cace use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents: 35385
diff changeset
   402
                                         (Const (maybe_name, T1 --> T1))))
42d39948cace use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents: 35385
diff changeset
   403
      in
42d39948cace use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents: 35385
diff changeset
   404
        if forall (fn (_, t) => t <> @{const True} andalso t <> @{const False})
42d39948cace use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents: 35385
diff changeset
   405
                  tps then
42d39948cace use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents: 35385
diff changeset
   406
          Const (unknown, T1 --> T2)
42d39948cace use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents: 35385
diff changeset
   407
        else
42d39948cace use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents: 35385
diff changeset
   408
          aux tps
42d39948cace use SAT4J for "Tests_Nits.thy" for safety (this should solve the Isatest failures) + minor changes
blanchet
parents: 35385
diff changeset
   409
      end
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   410
    fun make_map maybe_opt T1 T2 T2' =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   411
      let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   412
        val update_const = Const (@{const_name fun_upd},
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   413
                                  (T1 --> T2) --> T1 --> T2 --> T1 --> T2)
35402
115a5a95710a clarified @{const_name} vs. @{const_abbrev};
wenzelm
parents: 35388
diff changeset
   414
        fun aux' [] = Const (@{const_abbrev Map.empty}, T1 --> T2)
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   415
          | aux' ((t1, t2) :: tps) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   416
            (case t2 of
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   417
               Const (@{const_name None}, _) => aux' tps
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   418
             | _ => update_const $ aux' tps $ t1 $ t2)
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   419
        fun aux tps =
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   420
          if maybe_opt andalso not (is_complete_type datatypes false T1) then
37261
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
   421
            update_const $ aux' tps $ Const (unrep (), T1)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   422
            $ (Const (@{const_name Some}, T2' --> T2) $ Const (unknown, T2'))
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   423
          else
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   424
            aux' tps
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   425
      in aux end
34982
7b8c366e34a2 added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents: 34974
diff changeset
   426
    fun polish_funs Ts t =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   427
      (case fastype_of1 (Ts, t) of
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   428
         Type (@{type_name fun}, [T1, T2]) =>
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   429
         if is_plain_fun t then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   430
           case T2 of
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   431
             @{typ bool} =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   432
             let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   433
               val (maybe_opt, ts_pair) =
34982
7b8c366e34a2 added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents: 34974
diff changeset
   434
                 dest_plain_fun t ||> pairself (map (polish_funs Ts))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   435
             in
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   436
               make_set maybe_opt T1 T2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   437
                        (sort_wrt (truth_const_sort_key o snd) (op ~~ ts_pair))
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   438
             end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   439
           | Type (@{type_name option}, [T2']) =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   440
             let
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   441
               val (maybe_opt, ts_pair) =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   442
                 dest_plain_fun t ||> pairself (map (polish_funs Ts))
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   443
             in make_map maybe_opt T1 T2 T2' (rev (op ~~ ts_pair)) end
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   444
           | _ => raise SAME ()
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   445
         else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   446
           raise SAME ()
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   447
       | _ => raise SAME ())
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   448
      handle SAME () =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   449
             case t of
34982
7b8c366e34a2 added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents: 34974
diff changeset
   450
               (t1 as Const (@{const_name fun_upd}, _) $ t11 $ _)
7b8c366e34a2 added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents: 34974
diff changeset
   451
               $ (t2 as Const (s, _)) =>
7b8c366e34a2 added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents: 34974
diff changeset
   452
               if s = unknown then polish_funs Ts t11
7b8c366e34a2 added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents: 34974
diff changeset
   453
               else polish_funs Ts t1 $ polish_funs Ts t2
7b8c366e34a2 added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents: 34974
diff changeset
   454
             | t1 $ t2 => polish_funs Ts t1 $ polish_funs Ts t2
7b8c366e34a2 added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents: 34974
diff changeset
   455
             | Abs (s, T, t') => Abs (s, T, polish_funs (T :: Ts) t')
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   456
             | Const (s, Type (@{type_name fun}, [T1, T2])) =>
34982
7b8c366e34a2 added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents: 34974
diff changeset
   457
               if s = opt_flag orelse s = non_opt_flag then
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   458
                 Abs ("x", T1,
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   459
                      Const (if is_complete_type datatypes false T1 then
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   460
                               irrelevant
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   461
                             else
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   462
                               unknown, T2))
34982
7b8c366e34a2 added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents: 34974
diff changeset
   463
               else
7b8c366e34a2 added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents: 34974
diff changeset
   464
                 t
7b8c366e34a2 added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents: 34974
diff changeset
   465
             | t => t
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   466
    fun make_fun maybe_opt T1 T2 T' ts1 ts2 =
34126
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 34124
diff changeset
   467
      ts1 ~~ ts2 |> sort (nice_term_ord o pairself fst)
34998
5e492a862b34 four changes to Nitpick:
blanchet
parents: 34982
diff changeset
   468
                 |> make_plain_fun maybe_opt T1 T2
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   469
                 |> unarize_unbox_etc_term
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   470
                 |> typecast_fun (uniterize_unarize_unbox_etc_type T')
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   471
                                 (uniterize_unarize_unbox_etc_type T1)
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   472
                                 (uniterize_unarize_unbox_etc_type T2)
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   473
    fun term_for_atom seen (T as Type (@{type_name fun}, [T1, T2])) T' j _ =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   474
        let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   475
          val k1 = card_of_type card_assigns T1
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   476
          val k2 = card_of_type card_assigns T2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   477
        in
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   478
          term_for_rep true seen T T' (Vect (k1, Atom (k2, 0)))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   479
                       [nth_combination (replicate k1 (k2, 0)) j]
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   480
          handle General.Subscript =>
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33202
diff changeset
   481
                 raise ARG ("Nitpick_Model.reconstruct_term.term_for_atom",
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   482
                            signed_string_of_int j ^ " for " ^
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   483
                            string_for_rep (Vect (k1, Atom (k2, 0))))
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   484
        end
38190
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 38189
diff changeset
   485
      | term_for_atom seen (Type (@{type_name prod}, [T1, T2])) _ j k =
35075
888802be2019 handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents: 35070
diff changeset
   486
        let
888802be2019 handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents: 35070
diff changeset
   487
          val k1 = card_of_type card_assigns T1
888802be2019 handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents: 35070
diff changeset
   488
          val k2 = k div k1
888802be2019 handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents: 35070
diff changeset
   489
        in
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   490
          list_comb (HOLogic.pair_const T1 T2,
35075
888802be2019 handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents: 35070
diff changeset
   491
                     map3 (fn T => term_for_atom seen T T) [T1, T2]
888802be2019 handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents: 35070
diff changeset
   492
                          [j div k2, j mod k2] [k1, k2]) (* ### k2 or k1? FIXME *)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   493
        end
35075
888802be2019 handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents: 35070
diff changeset
   494
      | term_for_atom seen @{typ prop} _ j k =
888802be2019 handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents: 35070
diff changeset
   495
        HOLogic.mk_Trueprop (term_for_atom seen bool_T bool_T j k)
888802be2019 handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents: 35070
diff changeset
   496
      | term_for_atom _ @{typ bool} _ j _ =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   497
        if j = 0 then @{const False} else @{const True}
35075
888802be2019 handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents: 35070
diff changeset
   498
      | term_for_atom seen T _ j k =
35220
2bcdae5f4fdb added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents: 35190
diff changeset
   499
        if T = nat_T andalso is_standard_datatype thy stds nat_T then
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   500
          HOLogic.mk_number nat_T j
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   501
        else if T = int_T then
35075
888802be2019 handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents: 35070
diff changeset
   502
          HOLogic.mk_number int_T (int_for_atom (k, 0) j)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   503
        else if is_fp_iterator_type T then
35075
888802be2019 handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents: 35070
diff changeset
   504
          HOLogic.mk_number nat_T (k - j - 1)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   505
        else if T = @{typ bisim_iterator} then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   506
          HOLogic.mk_number nat_T j
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   507
        else case datatype_spec datatypes T of
37260
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   508
          NONE => nth_atom thy atomss pool for_auto T j
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   509
        | SOME {deep = false, ...} => nth_atom thy atomss pool for_auto T j
35179
4b198af5beb5 redo Nitpick's nonstandard values as cyclic values (instead of additional constructors)
blanchet
parents: 35177
diff changeset
   510
        | SOME {co, standard, constrs, ...} =>
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   511
          let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   512
            fun tuples_for_const (s, T) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   513
              tuple_list_for_name rel_table bounds (ConstName (s, T, Any))
35180
c57dba973391 more work on Nitpick's support for nonstandard models + fix in model reconstruction
blanchet
parents: 35179
diff changeset
   514
            fun cyclic_atom () =
37261
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
   515
              nth_atom thy atomss pool for_auto (Type (cyclic_type_name (), []))
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
   516
                       j
37260
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   517
            fun cyclic_var () =
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   518
              Var ((nth_atom_name thy atomss pool "" T j, 0), T)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   519
            val discr_jsss = map (tuples_for_const o discr_for_constr o #const)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   520
                                 constrs
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   521
            val real_j = j + offset_of_type ofs T
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   522
            val constr_x as (constr_s, constr_T) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   523
              get_first (fn (jss, {const, ...}) =>
34121
5e831d805118 get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents: 33982
diff changeset
   524
                            if member (op =) jss [real_j] then SOME const
5e831d805118 get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents: 33982
diff changeset
   525
                            else NONE)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   526
                        (discr_jsss ~~ constrs) |> the
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   527
            val arg_Ts = curried_binder_types constr_T
35190
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35188
diff changeset
   528
            val sel_xs =
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35188
diff changeset
   529
              map (binarized_and_boxed_nth_sel_for_constr hol_ctxt binarize
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35188
diff changeset
   530
                                                          constr_x)
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35188
diff changeset
   531
                  (index_seq 0 (length arg_Ts))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   532
            val sel_Rs =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   533
              map (fn x => get_first
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   534
                               (fn ConstName (s', T', R) =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   535
                                   if (s', T') = x then SOME R else NONE
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33202
diff changeset
   536
                                 | u => raise NUT ("Nitpick_Model.reconstruct_\
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   537
                                                   \term.term_for_atom", [u]))
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   538
                               sel_names |> the) sel_xs
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   539
            val arg_Rs = map (snd o dest_Func) sel_Rs
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   540
            val sel_jsss = map tuples_for_const sel_xs
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   541
            val arg_jsss =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   542
              map (map_filter (fn js => if hd js = real_j then SOME (tl js)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   543
                                        else NONE)) sel_jsss
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   544
            val uncur_arg_Ts = binder_types constr_T
35179
4b198af5beb5 redo Nitpick's nonstandard values as cyclic values (instead of additional constructors)
blanchet
parents: 35177
diff changeset
   545
            val maybe_cyclic = co orelse not standard
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   546
          in
35180
c57dba973391 more work on Nitpick's support for nonstandard models + fix in model reconstruction
blanchet
parents: 35179
diff changeset
   547
            if maybe_cyclic andalso not (null seen) andalso
35188
8c70a34931b1 improve Nitpick's "Datatypes" rendering for elements containing cycles
blanchet
parents: 35180
diff changeset
   548
               member (op =) (seen |> unfold ? (fst o split_last)) (T, j) then
35180
c57dba973391 more work on Nitpick's support for nonstandard models + fix in model reconstruction
blanchet
parents: 35179
diff changeset
   549
              cyclic_var ()
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   550
            else if constr_s = @{const_name Word} then
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   551
              HOLogic.mk_number
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   552
                  (if T = @{typ "unsigned_bit word"} then nat_T else int_T)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 34123
diff changeset
   553
                  (value_of_bits (the_single arg_jsss))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   554
            else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   555
              let
35179
4b198af5beb5 redo Nitpick's nonstandard values as cyclic values (instead of additional constructors)
blanchet
parents: 35177
diff changeset
   556
                val seen = seen |> maybe_cyclic ? cons (T, j)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   557
                val ts =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   558
                  if length arg_Ts = 0 then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   559
                    []
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   560
                  else
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   561
                    map3 (fn Ts =>
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   562
                             term_for_rep (constr_s <> @{const_name FinFun})
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   563
                                          seen Ts Ts) arg_Ts arg_Rs arg_jsss
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   564
                    |> mk_tuple (HOLogic.mk_tupleT uncur_arg_Ts)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   565
                    |> dest_n_tuple (length uncur_arg_Ts)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   566
                val t =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   567
                  if constr_s = @{const_name Abs_Frac} then
35711
548d3f16404b added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents: 35665
diff changeset
   568
                    case ts of
548d3f16404b added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents: 35665
diff changeset
   569
                      [Const (@{const_name Pair}, _) $ t1 $ t2] =>
548d3f16404b added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents: 35665
diff changeset
   570
                      frac_from_term_pair (body_type T) t1 t2
548d3f16404b added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents: 35665
diff changeset
   571
                    | _ => raise TERM ("Nitpick_Model.reconstruct_term.\
548d3f16404b added term postprocessor to Nitpick, to provide custom syntax for typedefs
blanchet
parents: 35665
diff changeset
   572
                                       \term_for_atom (Abs_Frac)", ts)
34936
c4f04bee79f3 some work on Nitpick's support for quotient types;
blanchet
parents: 34126
diff changeset
   573
                  else if not for_auto andalso
37256
0dca1ec52999 thread along context instead of theory for typedef lookup
blanchet
parents: 37170
diff changeset
   574
                          (is_abs_fun ctxt constr_x orelse
34936
c4f04bee79f3 some work on Nitpick's support for quotient types;
blanchet
parents: 34126
diff changeset
   575
                           constr_s = @{const_name Quot}) then
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   576
                    Const (abs_name, constr_T) $ the_single ts
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   577
                  else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   578
                    list_comb (Const constr_x, ts)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   579
              in
35179
4b198af5beb5 redo Nitpick's nonstandard values as cyclic values (instead of additional constructors)
blanchet
parents: 35177
diff changeset
   580
                if maybe_cyclic then
35180
c57dba973391 more work on Nitpick's support for nonstandard models + fix in model reconstruction
blanchet
parents: 35179
diff changeset
   581
                  let val var = cyclic_var () in
35188
8c70a34931b1 improve Nitpick's "Datatypes" rendering for elements containing cycles
blanchet
parents: 35180
diff changeset
   582
                    if unfold andalso not standard andalso
35180
c57dba973391 more work on Nitpick's support for nonstandard models + fix in model reconstruction
blanchet
parents: 35179
diff changeset
   583
                       length seen = 1 andalso
37261
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
   584
                       exists_subterm
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
   585
                           (fn Const (s, _) =>
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
   586
                               String.isPrefix (cyclic_const_prefix ()) s
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
   587
                             | t' => t' = var) t then
35188
8c70a34931b1 improve Nitpick's "Datatypes" rendering for elements containing cycles
blanchet
parents: 35180
diff changeset
   588
                      subst_atomic [(var, cyclic_atom ())] t
35180
c57dba973391 more work on Nitpick's support for nonstandard models + fix in model reconstruction
blanchet
parents: 35179
diff changeset
   589
                    else if exists_subterm (curry (op =) var) t then
35179
4b198af5beb5 redo Nitpick's nonstandard values as cyclic values (instead of additional constructors)
blanchet
parents: 35177
diff changeset
   590
                      if co then
4b198af5beb5 redo Nitpick's nonstandard values as cyclic values (instead of additional constructors)
blanchet
parents: 35177
diff changeset
   591
                        Const (@{const_name The}, (T --> bool_T) --> T)
37261
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
   592
                        $ Abs (cyclic_co_val_name (), T,
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38284
diff changeset
   593
                               Const (@{const_name HOL.eq}, T --> T --> bool_T)
35180
c57dba973391 more work on Nitpick's support for nonstandard models + fix in model reconstruction
blanchet
parents: 35179
diff changeset
   594
                               $ Bound 0 $ abstract_over (var, t))
35179
4b198af5beb5 redo Nitpick's nonstandard values as cyclic values (instead of additional constructors)
blanchet
parents: 35177
diff changeset
   595
                      else
35180
c57dba973391 more work on Nitpick's support for nonstandard models + fix in model reconstruction
blanchet
parents: 35179
diff changeset
   596
                        cyclic_atom ()
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   597
                    else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   598
                      t
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   599
                  end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   600
                else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   601
                  t
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   602
              end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   603
          end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   604
    and term_for_vect seen k R T1 T2 T' js =
35075
888802be2019 handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents: 35070
diff changeset
   605
      make_fun true T1 T2 T'
888802be2019 handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents: 35070
diff changeset
   606
               (map (fn j => term_for_atom seen T1 T1 j k) (index_seq 0 k))
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   607
               (map (term_for_rep true seen T2 T2 R o single)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   608
                    (batch_list (arity_of_rep R) js))
38190
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 38189
diff changeset
   609
    and term_for_rep _ seen T T' (R as Atom (k, j0)) [[j]] =
35075
888802be2019 handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents: 35070
diff changeset
   610
        if j >= j0 andalso j < j0 + k then term_for_atom seen T T' (j - j0) k
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33202
diff changeset
   611
        else raise REP ("Nitpick_Model.reconstruct_term.term_for_rep", [R])
38190
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 38189
diff changeset
   612
      | term_for_rep _ seen (Type (@{type_name prod}, [T1, T2])) _
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   613
                     (Struct [R1, R2]) [js] =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   614
        let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   615
          val arity1 = arity_of_rep R1
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   616
          val (js1, js2) = chop arity1 js
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   617
        in
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   618
          list_comb (HOLogic.pair_const T1 T2,
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   619
                     map3 (fn T => term_for_rep true seen T T) [T1, T2] [R1, R2]
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   620
                          [[js1], [js2]])
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   621
        end
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   622
      | term_for_rep _ seen (Type (@{type_name fun}, [T1, T2])) T'
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   623
                     (Vect (k, R')) [js] =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   624
        term_for_vect seen k R' T1 T2 T' js
37170
38ba15040455 Nitpick: show "..." in datatype values (e.g., [{0::nat, ...}]), since these are really equivalence classes
blanchet
parents: 36607
diff changeset
   625
      | term_for_rep maybe_opt seen (Type (@{type_name fun}, [T1, T2])) T'
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   626
                     (Func (R1, Formula Neut)) jss =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   627
        let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   628
          val jss1 = all_combinations_for_rep R1
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   629
          val ts1 = map (term_for_rep true seen T1 T1 R1 o single) jss1
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   630
          val ts2 =
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   631
            map (fn js => term_for_rep true seen T2 T2 (Atom (2, 0))
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35280
diff changeset
   632
                                       [[int_from_bool (member (op =) jss js)]])
34121
5e831d805118 get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents: 33982
diff changeset
   633
                jss1
37170
38ba15040455 Nitpick: show "..." in datatype values (e.g., [{0::nat, ...}]), since these are really equivalence classes
blanchet
parents: 36607
diff changeset
   634
        in make_fun maybe_opt T1 T2 T' ts1 ts2 end
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   635
      | term_for_rep maybe_opt seen (Type (@{type_name fun}, [T1, T2])) T'
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   636
                     (Func (R1, R2)) jss =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   637
        let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   638
          val arity1 = arity_of_rep R1
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   639
          val jss1 = all_combinations_for_rep R1
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   640
          val ts1 = map (term_for_rep false seen T1 T1 R1 o single) jss1
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   641
          val grouped_jss2 = AList.group (op =) (map (chop arity1) jss)
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   642
          val ts2 = map (term_for_rep false seen T2 T2 R2 o the_default []
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   643
                         o AList.lookup (op =) grouped_jss2) jss1
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   644
        in make_fun maybe_opt T1 T2 T' ts1 ts2 end
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   645
      | term_for_rep _ seen T T' (Opt R) jss =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   646
        if null jss then Const (unknown, T)
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   647
        else term_for_rep true seen T T' R jss
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   648
      | term_for_rep _ _ T _ R jss =
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33202
diff changeset
   649
        raise ARG ("Nitpick_Model.reconstruct_term.term_for_rep",
37260
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   650
                   Syntax.string_of_typ ctxt T ^ " " ^ string_for_rep R ^ " " ^
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   651
                   string_of_int (length jss))
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   652
  in
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   653
    postprocess_subterms [] o polish_funs [] o unarize_unbox_etc_term
37262
c0fe8fa35771 don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents: 37261
diff changeset
   654
    oooo term_for_rep maybe_opt []
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   655
  end
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   656
35718
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   657
(** Constant postprocessing **)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   658
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   659
fun dest_n_tuple_type 1 T = [T]
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   660
  | dest_n_tuple_type n (Type (_, [T1, T2])) =
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   661
    T1 :: dest_n_tuple_type (n - 1) T2
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   662
  | dest_n_tuple_type _ T =
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   663
    raise TYPE ("Nitpick_Model.dest_n_tuple_type", [T], [])
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   664
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   665
fun const_format thy def_table (x as (s, T)) =
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   666
  if String.isPrefix unrolled_prefix s then
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   667
    const_format thy def_table (original_name s, range_type T)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   668
  else if String.isPrefix skolem_prefix s then
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   669
    let
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   670
      val k = unprefix skolem_prefix s
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   671
              |> strip_first_name_sep |> fst |> space_explode "@"
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   672
              |> hd |> Int.fromString |> the
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   673
    in [k, num_binder_types T - k] end
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   674
  else if original_name s <> s then
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   675
    [num_binder_types T]
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   676
  else case def_of_const thy def_table x of
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   677
    SOME t' => if fixpoint_kind_of_rhs t' <> NoFp then
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   678
                 let val k = length (strip_abs_vars t') in
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   679
                   [k, num_binder_types T - k]
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   680
                 end
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   681
               else
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   682
                 [num_binder_types T]
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   683
  | NONE => [num_binder_types T]
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   684
fun intersect_formats _ [] = []
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   685
  | intersect_formats [] _ = []
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   686
  | intersect_formats ks1 ks2 =
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   687
    let val ((ks1', k1), (ks2', k2)) = pairself split_last (ks1, ks2) in
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   688
      intersect_formats (ks1' @ (if k1 > k2 then [k1 - k2] else []))
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   689
                        (ks2' @ (if k2 > k1 then [k2 - k1] else [])) @
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   690
      [Int.min (k1, k2)]
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   691
    end
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   692
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   693
fun lookup_format thy def_table formats t =
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   694
  case AList.lookup (fn (SOME x, SOME y) =>
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   695
                        (term_match thy) (x, y) | _ => false)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   696
                    formats (SOME t) of
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   697
    SOME format => format
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   698
  | NONE => let val format = the (AList.lookup (op =) formats NONE) in
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   699
              case t of
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   700
                Const x => intersect_formats format
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   701
                                             (const_format thy def_table x)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   702
              | _ => format
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   703
            end
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   704
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   705
fun format_type default_format format T =
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   706
  let
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   707
    val T = uniterize_unarize_unbox_etc_type T
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   708
    val format = format |> filter (curry (op <) 0)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   709
  in
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   710
    if forall (curry (op =) 1) format then
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   711
      T
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   712
    else
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   713
      let
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   714
        val (binder_Ts, body_T) = strip_type T
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   715
        val batched =
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   716
          binder_Ts
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   717
          |> map (format_type default_format default_format)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   718
          |> rev |> chunk_list_unevenly (rev format)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   719
          |> map (HOLogic.mk_tupleT o rev)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   720
      in List.foldl (op -->) body_T batched end
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   721
  end
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   722
fun format_term_type thy def_table formats t =
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   723
  format_type (the (AList.lookup (op =) formats NONE))
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   724
              (lookup_format thy def_table formats t) (fastype_of t)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   725
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   726
fun repair_special_format js m format =
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   727
  m - 1 downto 0 |> chunk_list_unevenly (rev format)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   728
                 |> map (rev o filter_out (member (op =) js))
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   729
                 |> filter_out null |> map length |> rev
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   730
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   731
fun user_friendly_const ({thy, evals, def_table, skolems, special_funs, ...}
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   732
                         : hol_context) (base_name, step_name) formats =
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   733
  let
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   734
    val default_format = the (AList.lookup (op =) formats NONE)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   735
    fun do_const (x as (s, T)) =
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   736
      (if String.isPrefix special_prefix s then
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   737
         let
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   738
           val do_term = map_aterms (fn Const x => fst (do_const x) | t' => t')
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   739
           val (x' as (_, T'), js, ts) =
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   740
             AList.find (op =) (!special_funs) (s, unarize_unbox_etc_type T)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   741
             |> the_single
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   742
           val max_j = List.last js
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   743
           val Ts = List.take (binder_types T', max_j + 1)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   744
           val missing_js = filter_out (member (op =) js) (0 upto max_j)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   745
           val missing_Ts = filter_indices missing_js Ts
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   746
           fun nth_missing_var n =
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   747
             ((arg_var_prefix ^ nat_subscript (n + 1), 0), nth missing_Ts n)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   748
           val missing_vars = map nth_missing_var (0 upto length missing_js - 1)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   749
           val vars = special_bounds ts @ missing_vars
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   750
           val ts' =
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   751
             map (fn j =>
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   752
                     case AList.lookup (op =) (js ~~ ts) j of
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   753
                       SOME t => do_term t
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   754
                     | NONE =>
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   755
                       Var (nth missing_vars
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   756
                                (find_index (curry (op =) j) missing_js)))
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   757
                 (0 upto max_j)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   758
           val t = do_const x' |> fst
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   759
           val format =
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   760
             case AList.lookup (fn (SOME t1, SOME t2) => term_match thy (t1, t2)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   761
                                 | _ => false) formats (SOME t) of
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   762
               SOME format =>
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   763
               repair_special_format js (num_binder_types T') format
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   764
             | NONE =>
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   765
               const_format thy def_table x'
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   766
               |> repair_special_format js (num_binder_types T')
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   767
               |> intersect_formats default_format
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   768
         in
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   769
           (list_comb (t, ts') |> fold_rev abs_var vars,
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   770
            format_type default_format format T)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   771
         end
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   772
       else if String.isPrefix uncurry_prefix s then
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   773
         let
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   774
           val (ss, s') = unprefix uncurry_prefix s
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   775
                          |> strip_first_name_sep |>> space_explode "@"
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   776
         in
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   777
           if String.isPrefix step_prefix s' then
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   778
             do_const (s', T)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   779
           else
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   780
             let
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   781
               val k = the (Int.fromString (hd ss))
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   782
               val j = the (Int.fromString (List.last ss))
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   783
               val (before_Ts, (tuple_T, rest_T)) =
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   784
                 strip_n_binders j T ||> (strip_n_binders 1 #>> hd)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   785
               val T' = before_Ts ---> dest_n_tuple_type k tuple_T ---> rest_T
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   786
             in do_const (s', T') end
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   787
         end
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   788
       else if String.isPrefix unrolled_prefix s then
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   789
         let val t = Const (original_name s, range_type T) in
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   790
           (lambda (Free (iter_var_prefix, nat_T)) t,
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   791
            format_type default_format
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   792
                        (lookup_format thy def_table formats t) T)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   793
         end
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   794
       else if String.isPrefix base_prefix s then
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   795
         (Const (base_name, T --> T) $ Const (unprefix base_prefix s, T),
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   796
          format_type default_format default_format T)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   797
       else if String.isPrefix step_prefix s then
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   798
         (Const (step_name, T --> T) $ Const (unprefix step_prefix s, T),
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   799
          format_type default_format default_format T)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   800
       else if String.isPrefix quot_normal_prefix s then
38207
792b78e355e7 added support for "Abs_" and "Rep_" functions on quotient types
blanchet
parents: 38190
diff changeset
   801
         let val t = Const (nitpick_prefix ^ "quotient normal form", T) in
35718
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   802
           (t, format_term_type thy def_table formats t)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   803
         end
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   804
       else if String.isPrefix skolem_prefix s then
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   805
         let
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   806
           val ss = the (AList.lookup (op =) (!skolems) s)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   807
           val (Ts, Ts') = chop (length ss) (binder_types T)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   808
           val frees = map Free (ss ~~ Ts)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   809
           val s' = original_name s
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   810
         in
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   811
           (fold lambda frees (Const (s', Ts' ---> T)),
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   812
            format_type default_format
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   813
                        (lookup_format thy def_table formats (Const x)) T)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   814
         end
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   815
       else if String.isPrefix eval_prefix s then
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   816
         let
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   817
           val t = nth evals (the (Int.fromString (unprefix eval_prefix s)))
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   818
         in (t, format_term_type thy def_table formats t) end
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   819
       else if s = @{const_name undefined_fast_The} then
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   820
         (Const (nitpick_prefix ^ "The fallback", T),
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   821
          format_type default_format
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   822
                      (lookup_format thy def_table formats
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   823
                           (Const (@{const_name The}, (T --> bool_T) --> T))) T)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   824
       else if s = @{const_name undefined_fast_Eps} then
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   825
         (Const (nitpick_prefix ^ "Eps fallback", T),
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   826
          format_type default_format
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   827
                      (lookup_format thy def_table formats
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   828
                           (Const (@{const_name Eps}, (T --> bool_T) --> T))) T)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   829
       else
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   830
         let val t = Const (original_name s, T) in
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   831
           (t, format_term_type thy def_table formats t)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   832
         end)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   833
      |>> map_types uniterize_unarize_unbox_etc_type
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   834
      |>> shorten_names_in_term |>> Term.map_abs_vars shortest_name
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   835
  in do_const end
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   836
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   837
fun assign_operator_for_const (s, T) =
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   838
  if String.isPrefix ubfp_prefix s then
37261
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
   839
    if is_fun_type T then xsym "\<subseteq>" "<=" ()
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
   840
    else xsym "\<le>" "<=" ()
35718
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   841
  else if String.isPrefix lbfp_prefix s then
37261
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
   842
    if is_fun_type T then xsym "\<supseteq>" ">=" ()
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
   843
    else xsym "\<ge>" ">=" ()
35718
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   844
  else if original_name s <> s then
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   845
    assign_operator_for_const (strip_first_name_sep s |> snd, T)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   846
  else
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   847
    "="
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   848
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   849
(** Model reconstruction **)
eee1a5e0d334 moved some Nitpick code around
blanchet
parents: 35712
diff changeset
   850
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   851
fun unfold_outer_the_binders (t as Const (@{const_name The}, _)
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38284
diff changeset
   852
                                   $ Abs (s, T, Const (@{const_name HOL.eq}, _)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   853
                                                $ Bound 0 $ t')) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   854
    betapply (Abs (s, T, t'), t) |> unfold_outer_the_binders
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   855
  | unfold_outer_the_binders t = t
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   856
fun bisimilar_values _ 0 _ = true
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   857
  | bisimilar_values coTs max_depth (t1, t2) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   858
    let val T = fastype_of t1 in
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   859
      if exists_subtype (member (op =) coTs) T then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   860
        let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   861
          val ((head1, args1), (head2, args2)) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   862
            pairself (strip_comb o unfold_outer_the_binders) (t1, t2)
34121
5e831d805118 get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents: 33982
diff changeset
   863
          val max_depth = max_depth - (if member (op =) coTs T then 1 else 0)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   864
        in
34936
c4f04bee79f3 some work on Nitpick's support for quotient types;
blanchet
parents: 34126
diff changeset
   865
          head1 = head2 andalso
c4f04bee79f3 some work on Nitpick's support for quotient types;
blanchet
parents: 34126
diff changeset
   866
          forall (bisimilar_values coTs max_depth) (args1 ~~ args2)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   867
        end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   868
      else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   869
        t1 = t2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   870
    end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   871
36390
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36389
diff changeset
   872
fun reconstruct_hol_model {show_datatypes, show_consts}
35280
54ab4921f826 fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents: 35220
diff changeset
   873
        ({hol_ctxt = {thy, ctxt, max_bisim_depth, boxes, stds, wfs, user_axioms,
38209
3d1d928dce50 added "whack"
blanchet
parents: 38207
diff changeset
   874
                      debug, whacks, binary_ints, destroy_constrs, specialize,
39359
6f49c7fbb1b1 remove "fast_descs" option from Nitpick;
blanchet
parents: 39345
diff changeset
   875
                      star_linear_preds, tac_timeout, evals, case_names,
6f49c7fbb1b1 remove "fast_descs" option from Nitpick;
blanchet
parents: 39345
diff changeset
   876
                      def_table, nondef_table, user_nondefs, simp_table,
6f49c7fbb1b1 remove "fast_descs" option from Nitpick;
blanchet
parents: 39345
diff changeset
   877
                      psimp_table, choice_spec_table, intro_table,
36388
30f7ce76712d removed Nitpick's "uncurry" option
blanchet
parents: 36385
diff changeset
   878
                      ground_thm_table, ersatz_table, skolems, special_funs,
30f7ce76712d removed Nitpick's "uncurry" option
blanchet
parents: 36385
diff changeset
   879
                      unrolled_preds, wf_cache, constr_cache},
35190
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35188
diff changeset
   880
         binarize, card_assigns, bits, bisim_depth, datatypes, ofs} : scope)
38170
d74b66ec02ce handle free variables even more gracefully;
blanchet
parents: 38126
diff changeset
   881
        formats atomss real_frees pseudo_frees free_names sel_names nonsel_names
d74b66ec02ce handle free variables even more gracefully;
blanchet
parents: 38126
diff changeset
   882
        rel_table bounds =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   883
  let
35180
c57dba973391 more work on Nitpick's support for nonstandard models + fix in model reconstruction
blanchet
parents: 35179
diff changeset
   884
    val pool = Unsynchronized.ref []
35280
54ab4921f826 fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents: 35220
diff changeset
   885
    val (wacky_names as (_, base_step_names), ctxt) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   886
      add_wacky_syntax ctxt
35070
96136eb6218f split "nitpick_hol.ML" into two files to make it more manageable;
blanchet
parents: 34998
diff changeset
   887
    val hol_ctxt =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   888
      {thy = thy, ctxt = ctxt, max_bisim_depth = max_bisim_depth, boxes = 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: 34974
diff changeset
   889
       stds = stds, wfs = wfs, user_axioms = user_axioms, debug = debug,
38209
3d1d928dce50 added "whack"
blanchet
parents: 38207
diff changeset
   890
       whacks = whacks, binary_ints = binary_ints,
3d1d928dce50 added "whack"
blanchet
parents: 38207
diff changeset
   891
       destroy_constrs = destroy_constrs, specialize = specialize,
39359
6f49c7fbb1b1 remove "fast_descs" option from Nitpick;
blanchet
parents: 39345
diff changeset
   892
       star_linear_preds = star_linear_preds, tac_timeout = tac_timeout,
6f49c7fbb1b1 remove "fast_descs" option from Nitpick;
blanchet
parents: 39345
diff changeset
   893
       evals = evals, case_names = case_names, def_table = def_table,
6f49c7fbb1b1 remove "fast_descs" option from Nitpick;
blanchet
parents: 39345
diff changeset
   894
       nondef_table = nondef_table, user_nondefs = user_nondefs,
6f49c7fbb1b1 remove "fast_descs" option from Nitpick;
blanchet
parents: 39345
diff changeset
   895
       simp_table = simp_table, psimp_table = psimp_table,
6f49c7fbb1b1 remove "fast_descs" option from Nitpick;
blanchet
parents: 39345
diff changeset
   896
       choice_spec_table = choice_spec_table, intro_table = intro_table,
6f49c7fbb1b1 remove "fast_descs" option from Nitpick;
blanchet
parents: 39345
diff changeset
   897
       ground_thm_table = ground_thm_table, ersatz_table = ersatz_table,
6f49c7fbb1b1 remove "fast_descs" option from Nitpick;
blanchet
parents: 39345
diff changeset
   898
       skolems = skolems, special_funs = special_funs,
6f49c7fbb1b1 remove "fast_descs" option from Nitpick;
blanchet
parents: 39345
diff changeset
   899
       unrolled_preds = unrolled_preds, wf_cache = wf_cache,
6f49c7fbb1b1 remove "fast_descs" option from Nitpick;
blanchet
parents: 39345
diff changeset
   900
       constr_cache = constr_cache}
36388
30f7ce76712d removed Nitpick's "uncurry" option
blanchet
parents: 36385
diff changeset
   901
    val scope =
30f7ce76712d removed Nitpick's "uncurry" option
blanchet
parents: 36385
diff changeset
   902
      {hol_ctxt = hol_ctxt, binarize = binarize, card_assigns = card_assigns,
30f7ce76712d removed Nitpick's "uncurry" option
blanchet
parents: 36385
diff changeset
   903
       bits = bits, bisim_depth = bisim_depth, datatypes = datatypes, ofs = ofs}
37262
c0fe8fa35771 don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents: 37261
diff changeset
   904
    fun term_for_rep maybe_opt unfold =
c0fe8fa35771 don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents: 37261
diff changeset
   905
      reconstruct_term maybe_opt unfold pool wacky_names scope atomss
c0fe8fa35771 don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents: 37261
diff changeset
   906
                       sel_names rel_table bounds
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   907
    val all_values =
37260
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   908
      all_values_of_type pool wacky_names scope atomss sel_names rel_table
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   909
                         bounds
38126
8031d099379a added manual symmetry breaking for datatypes
blanchet
parents: 37678
diff changeset
   910
    fun is_codatatype_wellformed (cos : datatype_spec list)
8031d099379a added manual symmetry breaking for datatypes
blanchet
parents: 37678
diff changeset
   911
                                 ({typ, card, ...} : datatype_spec) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   912
      let
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   913
        val ts = all_values card typ
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   914
        val max_depth = Integer.sum (map #card cos)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   915
      in
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   916
        forall (not o bisimilar_values (map #typ cos) max_depth)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   917
               (all_distinct_unordered_pairs_of ts)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   918
      end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   919
    fun pretty_for_assign name =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   920
      let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   921
        val (oper, (t1, T'), T) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   922
          case name of
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   923
            FreeName (s, T, _) =>
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   924
            let val t = Free (s, uniterize_unarize_unbox_etc_type T) in
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   925
              ("=", (t, format_term_type thy def_table formats t), T)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   926
            end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   927
          | ConstName (s, T, _) =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   928
            (assign_operator_for_const (s, T),
35280
54ab4921f826 fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents: 35220
diff changeset
   929
             user_friendly_const hol_ctxt base_step_names formats (s, T), T)
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33202
diff changeset
   930
          | _ => raise NUT ("Nitpick_Model.reconstruct_hol_model.\
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   931
                            \pretty_for_assign", [name])
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   932
        val t2 = if rep_of name = Any then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   933
                   Const (@{const_name undefined}, T')
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   934
                 else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   935
                   tuple_list_for_name rel_table bounds name
37262
c0fe8fa35771 don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents: 37261
diff changeset
   936
                   |> term_for_rep (not (is_fully_representable_set name)) false
c0fe8fa35771 don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents: 37261
diff changeset
   937
                                   T T' (rep_of name)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   938
      in
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   939
        Pretty.block (Pretty.breaks
39118
12f3788be67b turned show_all_types into proper configuration option;
wenzelm
parents: 38864
diff changeset
   940
            [Syntax.pretty_term (set_show_all_types ctxt) t1,
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   941
             Pretty.str oper, Syntax.pretty_term ctxt t2])
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   942
      end
38126
8031d099379a added manual symmetry breaking for datatypes
blanchet
parents: 37678
diff changeset
   943
    fun pretty_for_datatype ({typ, card, complete, ...} : datatype_spec) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   944
      Pretty.block (Pretty.breaks
38189
blanchet
parents: 38170
diff changeset
   945
          (pretty_for_type ctxt typ ::
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   946
           (case typ of
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   947
              Type (@{type_name fin_fun}, _) => [Pretty.str "[finite]"]
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   948
            | Type (@{type_name fun_box}, _) => [Pretty.str "[boxed]"]
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   949
            | Type (@{type_name pair_box}, _) => [Pretty.str "[boxed]"]
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   950
            | _ => []) @
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   951
           [Pretty.str "=",
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   952
            Pretty.enum "," "{" "}"
35712
77aa29bf14ee added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
blanchet
parents: 35711
diff changeset
   953
                (map (Syntax.pretty_term ctxt) (all_values card typ) @
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   954
                 (if fun_from_pair complete false then []
37261
8a89fd40ed0b honor xsymbols in Nitpick
blanchet
parents: 37260
diff changeset
   955
                  else [Pretty.str (unrep ())]))]))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   956
    fun integer_datatype T =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   957
      [{typ = T, card = card_of_type card_assigns T, co = false,
38126
8031d099379a added manual symmetry breaking for datatypes
blanchet
parents: 37678
diff changeset
   958
        standard = true, self_rec = true, complete = (false, false),
8031d099379a added manual symmetry breaking for datatypes
blanchet
parents: 37678
diff changeset
   959
        concrete = (true, true), deep = true, constrs = []}]
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33202
diff changeset
   960
      handle TYPE ("Nitpick_HOL.card_of_type", _, _) => []
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   961
    val (codatatypes, datatypes) =
34982
7b8c366e34a2 added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents: 34974
diff changeset
   962
      datatypes |> filter #deep |> List.partition #co
35220
2bcdae5f4fdb added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents: 35190
diff changeset
   963
                ||> append (integer_datatype int_T
2bcdae5f4fdb added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents: 35190
diff changeset
   964
                            |> is_standard_datatype thy stds nat_T
2bcdae5f4fdb added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents: 35190
diff changeset
   965
                               ? append (integer_datatype nat_T))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   966
    val block_of_datatypes =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   967
      if show_datatypes andalso not (null datatypes) then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   968
        [Pretty.big_list ("Datatype" ^ plural_s_for_list datatypes ^ ":")
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   969
                         (map pretty_for_datatype datatypes)]
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   970
      else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   971
        []
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   972
    val block_of_codatatypes =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   973
      if show_datatypes andalso not (null codatatypes) then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   974
        [Pretty.big_list ("Codatatype" ^ plural_s_for_list codatatypes ^ ":")
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   975
                         (map pretty_for_datatype codatatypes)]
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   976
      else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   977
        []
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   978
    fun block_of_names show title names =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   979
      if show andalso not (null names) then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   980
        Pretty.str (title ^ plural_s_for_list names ^ ":")
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   981
        :: map (Pretty.indent indent_size o pretty_for_assign)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   982
               (sort_wrt (original_name o nickname_of) names)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   983
      else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   984
        []
38170
d74b66ec02ce handle free variables even more gracefully;
blanchet
parents: 38126
diff changeset
   985
    fun free_name_for_term keep_all (x as (s, T)) =
d74b66ec02ce handle free variables even more gracefully;
blanchet
parents: 38126
diff changeset
   986
      case filter (curry (op =) x
d74b66ec02ce handle free variables even more gracefully;
blanchet
parents: 38126
diff changeset
   987
                   o pairf nickname_of (uniterize_unarize_unbox_etc_type
d74b66ec02ce handle free variables even more gracefully;
blanchet
parents: 38126
diff changeset
   988
                                        o type_of)) free_names of
d74b66ec02ce handle free variables even more gracefully;
blanchet
parents: 38126
diff changeset
   989
        [name] => SOME name
d74b66ec02ce handle free variables even more gracefully;
blanchet
parents: 38126
diff changeset
   990
      | [] => if keep_all then SOME (FreeName (s, T, Any)) else NONE
d74b66ec02ce handle free variables even more gracefully;
blanchet
parents: 38126
diff changeset
   991
      | _ => raise TERM ("Nitpick_Model.reconstruct_hol_model.\
d74b66ec02ce handle free variables even more gracefully;
blanchet
parents: 38126
diff changeset
   992
                         \free_name_for_term", [Const x])
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   993
    val (skolem_names, nonskolem_nonsel_names) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   994
      List.partition is_skolem_name nonsel_names
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   995
    val (eval_names, noneval_nonskolem_nonsel_names) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   996
      List.partition (String.isPrefix eval_prefix o nickname_of)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   997
                     nonskolem_nonsel_names
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   998
      ||> filter_out (member (op =) [@{const_name bisim},
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   999
                                     @{const_name bisim_iterator_max}]
34121
5e831d805118 get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents: 33982
diff changeset
  1000
                      o nickname_of)
38170
d74b66ec02ce handle free variables even more gracefully;
blanchet
parents: 38126
diff changeset
  1001
      ||> append (map_filter (free_name_for_term false) pseudo_frees)
d74b66ec02ce handle free variables even more gracefully;
blanchet
parents: 38126
diff changeset
  1002
    val real_free_names = map_filter (free_name_for_term true) real_frees
d74b66ec02ce handle free variables even more gracefully;
blanchet
parents: 38126
diff changeset
  1003
    val chunks = block_of_names true "Free variable" real_free_names @
36390
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36389
diff changeset
  1004
                 block_of_names true "Skolem constant" skolem_names @
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1005
                 block_of_names true "Evaluated term" eval_names @
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1006
                 block_of_datatypes @ block_of_codatatypes @
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1007
                 block_of_names show_consts "Constant"
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1008
                                noneval_nonskolem_nonsel_names
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1009
  in
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1010
    (Pretty.chunks (if null chunks then [Pretty.str "Empty assignment"]
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1011
                    else chunks),
34936
c4f04bee79f3 some work on Nitpick's support for quotient types;
blanchet
parents: 34126
diff changeset
  1012
     bisim_depth >= 0 orelse
c4f04bee79f3 some work on Nitpick's support for quotient types;
blanchet
parents: 34126
diff changeset
  1013
     forall (is_codatatype_wellformed codatatypes) codatatypes)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1014
  end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1015
37262
c0fe8fa35771 don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents: 37261
diff changeset
  1016
fun term_for_name pool scope atomss sel_names rel_table bounds name =
c0fe8fa35771 don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents: 37261
diff changeset
  1017
  let val T = type_of name in
c0fe8fa35771 don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents: 37261
diff changeset
  1018
    tuple_list_for_name rel_table bounds name
c0fe8fa35771 don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents: 37261
diff changeset
  1019
    |> reconstruct_term (not (is_fully_representable_set name)) false pool
c0fe8fa35771 don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents: 37261
diff changeset
  1020
                        (("", ""), ("", "")) scope atomss sel_names rel_table
c0fe8fa35771 don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents: 37261
diff changeset
  1021
                        bounds T T (rep_of name)
c0fe8fa35771 don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents: 37261
diff changeset
  1022
  end
c0fe8fa35771 don't show spurious "..." in Nitpick's output for free variables of set type (e.g., P (op +) example from Manual_Nits.thy); undoes parts of 38ba15040455, which was too aggressive
blanchet
parents: 37261
diff changeset
  1023
35280
54ab4921f826 fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents: 35220
diff changeset
  1024
fun prove_hol_model (scope as {hol_ctxt = {thy, ctxt, debug, ...},
34998
5e492a862b34 four changes to Nitpick:
blanchet
parents: 34982
diff changeset
  1025
                               card_assigns, ...})
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1026
                    auto_timeout free_names sel_names rel_table bounds prop =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1027
  let
35076
cc19e2aef17e added hotel key card example for Nitpick, and renumber atoms in Nitpick's output for increased readability
blanchet
parents: 35075
diff changeset
  1028
    val pool = Unsynchronized.ref []
37260
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
  1029
    val atomss = [(NONE, [])]
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1030
    fun free_type_assm (T, k) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1031
      let
37260
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
  1032
        fun atom j = nth_atom thy atomss pool true T j
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1033
        fun equation_for_atom j = HOLogic.eq_const T $ Bound 0 $ atom j
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1034
        val eqs = map equation_for_atom (index_seq 0 k)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1035
        val compreh_assm =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1036
          Const (@{const_name All}, (T --> bool_T) --> bool_T)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1037
              $ Abs ("x", T, foldl1 HOLogic.mk_disj eqs)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1038
        val distinct_assm = distinctness_formula T (map atom (index_seq 0 k))
34998
5e492a862b34 four changes to Nitpick:
blanchet
parents: 34982
diff changeset
  1039
      in s_conj (compreh_assm, distinct_assm) end
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1040
    fun free_name_assm name =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1041
      HOLogic.mk_eq (Free (nickname_of name, type_of name),
37260
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
  1042
                     term_for_name pool scope atomss sel_names rel_table bounds
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
  1043
                                   name)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1044
    val freeT_assms = map free_type_assm (filter (is_TFree o fst) card_assigns)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1045
    val model_assms = map free_name_assm free_names
34998
5e492a862b34 four changes to Nitpick:
blanchet
parents: 34982
diff changeset
  1046
    val assm = foldr1 s_conj (freeT_assms @ model_assms)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1047
    fun try_out negate =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1048
      let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1049
        val concl = (negate ? curry (op $) @{const Not})
35625
9c818cab0dd0 modernized structure Object_Logic;
wenzelm
parents: 35408
diff changeset
  1050
                    (Object_Logic.atomize_term thy prop)
34998
5e492a862b34 four changes to Nitpick:
blanchet
parents: 34982
diff changeset
  1051
        val prop = HOLogic.mk_Trueprop (HOLogic.mk_imp (assm, concl))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1052
                   |> map_types (map_type_tfree
34998
5e492a862b34 four changes to Nitpick:
blanchet
parents: 34982
diff changeset
  1053
                                     (fn (s, []) => TFree (s, HOLogic.typeS)
5e492a862b34 four changes to Nitpick:
blanchet
parents: 34982
diff changeset
  1054
                                       | x => TFree x))
5e492a862b34 four changes to Nitpick:
blanchet
parents: 34982
diff changeset
  1055
       val _ = if debug then
5e492a862b34 four changes to Nitpick:
blanchet
parents: 34982
diff changeset
  1056
                 priority ((if negate then "Genuineness" else "Spuriousness") ^
5e492a862b34 four changes to Nitpick:
blanchet
parents: 34982
diff changeset
  1057
                           " goal: " ^ Syntax.string_of_term ctxt prop ^ ".")
5e492a862b34 four changes to Nitpick:
blanchet
parents: 34982
diff changeset
  1058
               else
5e492a862b34 four changes to Nitpick:
blanchet
parents: 34982
diff changeset
  1059
                 ()
5e492a862b34 four changes to Nitpick:
blanchet
parents: 34982
diff changeset
  1060
        val goal = prop |> cterm_of thy |> Goal.init
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1061
      in
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1062
        (goal |> SINGLE (DETERM_TIMEOUT auto_timeout
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1063
                                        (auto_tac (clasimpset_of ctxt)))
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1064
              |> the |> Goal.finish ctxt; true)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1065
        handle THM _ => false
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1066
             | TimeLimit.TimeOut => false
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1067
      end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1068
  in
33705
947184dc75c9 removed a few global names in Nitpick (styp, nat_less, pairf)
blanchet
parents: 33580
diff changeset
  1069
    if try_out false then SOME true
947184dc75c9 removed a few global names in Nitpick (styp, nat_less, pairf)
blanchet
parents: 33580
diff changeset
  1070
    else if try_out true then SOME false
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1071
    else NONE
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1072
  end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1073
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1074
end;