src/HOL/Tools/Nitpick/nitpick_mono.ML
author blanchet
Mon, 06 Dec 2010 13:18:25 +0100
changeset 40985 8b870370c26f
parent 40722 441260986b63
child 40986 cfd91aa80701
permissions -rw-r--r--
started generalizing monotonicity code to accommodate new calculus
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33982
1ae222745c4a fixed paths in Nitpick's ML file headers
blanchet
parents: 33852
diff changeset
     1
(*  Title:      HOL/Tools/Nitpick/nitpick_mono.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: 34936
diff changeset
     3
    Copyright   2009, 2010
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
     4
35384
88dbcfe75c45 cosmetics
blanchet
parents: 35280
diff changeset
     5
Monotonicity inference for higher-order logic.
33192
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_MONO =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
     9
sig
35070
96136eb6218f split "nitpick_hol.ML" into two files to make it more manageable;
blanchet
parents: 34998
diff changeset
    10
  type hol_context = Nitpick_HOL.hol_context
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    11
38179
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
    12
  val trace : bool Unsynchronized.ref
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    13
  val formulas_monotonic :
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
    14
    hol_context -> bool -> typ -> term list * term list -> bool
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
    15
  val finitize_funs :
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
    16
    hol_context -> bool -> (typ option * bool option) list -> typ
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
    17
    -> term list * term list -> term list * term list
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    18
end;
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    19
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
    20
structure Nitpick_Mono : NITPICK_MONO =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    21
struct
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    22
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
    23
open Nitpick_Util
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
    24
open Nitpick_HOL
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    25
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
    26
datatype sign = Plus | Minus
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
    27
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    28
type var = int
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    29
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
    30
datatype annotation = Gen | New | Fls | Tru
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
    31
datatype annotation_atom = A of annotation | V of var
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    32
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
    33
type literal = var * annotation
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    34
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
    35
datatype mtyp =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
    36
  MAlpha |
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
    37
  MFun of mtyp * annotation_atom * mtyp |
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
    38
  MPair of mtyp * mtyp |
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
    39
  MType of string * mtyp list |
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
    40
  MRec of string * typ list
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    41
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
    42
datatype mterm =
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
    43
  MRaw of term * mtyp |
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
    44
  MAbs of string * typ * mtyp * annotation_atom * mterm |
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
    45
  MApp of mterm * mterm
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
    46
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
    47
type mdata =
35070
96136eb6218f split "nitpick_hol.ML" into two files to make it more manageable;
blanchet
parents: 34998
diff changeset
    48
  {hol_ctxt: hol_context,
35190
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35079
diff changeset
    49
   binarize: bool,
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    50
   alpha_T: typ,
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
    51
   no_harmless: bool,
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    52
   max_fresh: int Unsynchronized.ref,
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
    53
   datatype_mcache: ((string * typ list) * mtyp) list Unsynchronized.ref,
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
    54
   constr_mcache: (styp * mtyp) list Unsynchronized.ref}
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    55
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
    56
exception UNSOLVABLE of unit
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
    57
exception MTYPE of string * mtyp list * typ list
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
    58
exception MTERM of string * mterm list
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    59
38179
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
    60
val trace = Unsynchronized.ref false
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
    61
fun trace_msg msg = if !trace then tracing (msg ()) else ()
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    62
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
    63
fun string_for_sign Plus = "+"
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
    64
  | string_for_sign Minus = "-"
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
    65
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
    66
fun negate_sign Plus = Minus
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
    67
  | negate_sign Minus = Plus
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
    68
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    69
val string_for_var = signed_string_of_int
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    70
fun string_for_vars sep [] = "0\<^bsub>" ^ sep ^ "\<^esub>"
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    71
  | string_for_vars sep xs = space_implode sep (map string_for_var xs)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    72
fun subscript_string_for_vars sep xs =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    73
  if null xs then "" else "\<^bsub>" ^ string_for_vars sep xs ^ "\<^esub>"
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    74
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
    75
fun string_for_annotation Gen = "G"
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
    76
  | string_for_annotation New = "N"
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
    77
  | string_for_annotation Fls = "F"
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
    78
  | string_for_annotation Tru = "T"
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    79
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
    80
fun string_for_annotation_atom (A a) = string_for_annotation a
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
    81
  | string_for_annotation_atom (V x) = string_for_var x
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    82
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
    83
fun string_for_literal (x, a) =
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
    84
  string_for_var x ^ " = " ^ string_for_annotation a
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    85
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
    86
val bool_M = MType (@{type_name bool}, [])
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
    87
val dummy_M = MType (nitpick_prefix ^ "dummy", [])
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    88
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
    89
fun is_MRec (MRec _) = true
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
    90
  | is_MRec _ = false
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
    91
fun dest_MFun (MFun z) = z
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
    92
  | dest_MFun M = raise MTYPE ("Nitpick_Mono.dest_MFun", [M], [])
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    93
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    94
val no_prec = 100
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    95
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
    96
fun precedence_of_mtype (MFun _) = 1
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
    97
  | precedence_of_mtype (MPair _) = 2
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
    98
  | precedence_of_mtype _ = no_prec
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    99
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   100
val string_for_mtype =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   101
  let
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   102
    fun aux outer_prec M =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   103
      let
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   104
        val prec = precedence_of_mtype M
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   105
        val need_parens = (prec < outer_prec)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   106
      in
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   107
        (if need_parens then "(" else "") ^
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   108
        (if M = dummy_M then
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   109
           "_"
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   110
         else case M of
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   111
             MAlpha => "\<alpha>"
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   112
           | MFun (M1, a, M2) =>
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   113
             aux (prec + 1) M1 ^ " \<Rightarrow>\<^bsup>" ^
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   114
             string_for_annotation_atom a ^ "\<^esup> " ^ aux prec M2
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   115
           | MPair (M1, M2) => aux (prec + 1) M1 ^ " \<times> " ^ aux prec M2
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   116
           | MType (s, []) =>
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   117
             if s = @{type_name prop} orelse s = @{type_name bool} then "o"
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   118
             else s
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   119
           | MType (s, Ms) => "(" ^ commas (map (aux 0) Ms) ^ ") " ^ s
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   120
           | MRec (s, _) => "[" ^ s ^ "]") ^
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   121
        (if need_parens then ")" else "")
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   122
      end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   123
  in aux 0 end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   124
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   125
fun flatten_mtype (MPair (M1, M2)) = maps flatten_mtype [M1, M2]
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   126
  | flatten_mtype (MType (_, Ms)) = maps flatten_mtype Ms
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   127
  | flatten_mtype M = [M]
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   128
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   129
fun precedence_of_mterm (MRaw _) = no_prec
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   130
  | precedence_of_mterm (MAbs _) = 1
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   131
  | precedence_of_mterm (MApp _) = 2
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   132
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   133
fun string_for_mterm ctxt =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   134
  let
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   135
    fun mtype_annotation M = "\<^bsup>" ^ string_for_mtype M ^ "\<^esup>"
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   136
    fun aux outer_prec m =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   137
      let
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   138
        val prec = precedence_of_mterm m
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   139
        val need_parens = (prec < outer_prec)
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   140
      in
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   141
        (if need_parens then "(" else "") ^
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   142
        (case m of
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   143
           MRaw (t, M) => Syntax.string_of_term ctxt t ^ mtype_annotation M
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   144
         | MAbs (s, _, M, a, m) =>
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   145
           "\<lambda>" ^ s ^ mtype_annotation M ^ ".\<^bsup>" ^
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   146
           string_for_annotation_atom a ^ "\<^esup> " ^ aux prec m
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   147
         | MApp (m1, m2) => aux prec m1 ^ " " ^ aux (prec + 1) m2) ^
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   148
        (if need_parens then ")" else "")
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   149
      end
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   150
  in aux 0 end
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   151
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   152
fun mtype_of_mterm (MRaw (_, M)) = M
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   153
  | mtype_of_mterm (MAbs (_, _, M, a, m)) = MFun (M, a, mtype_of_mterm m)
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   154
  | mtype_of_mterm (MApp (m1, _)) =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   155
    case mtype_of_mterm m1 of
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   156
      MFun (_, _, M12) => M12
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   157
    | M1 => raise MTYPE ("Nitpick_Mono.mtype_of_mterm", [M1], [])
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   158
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   159
fun strip_mcomb (MApp (m1, m2)) = strip_mcomb m1 ||> (fn ms => append ms [m2])
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   160
  | strip_mcomb m = (m, [])
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   161
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   162
fun initial_mdata hol_ctxt binarize no_harmless alpha_T =
35190
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35079
diff changeset
   163
  ({hol_ctxt = hol_ctxt, binarize = binarize, alpha_T = alpha_T,
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   164
    no_harmless = no_harmless, max_fresh = Unsynchronized.ref 0,
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   165
    datatype_mcache = Unsynchronized.ref [],
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   166
    constr_mcache = Unsynchronized.ref []} : mdata)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   167
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   168
fun could_exist_alpha_subtype alpha_T (T as Type (_, Ts)) =
34936
c4f04bee79f3 some work on Nitpick's support for quotient types;
blanchet
parents: 34123
diff changeset
   169
    T = alpha_T orelse (not (is_fp_iterator_type T) andalso
c4f04bee79f3 some work on Nitpick's support for quotient types;
blanchet
parents: 34123
diff changeset
   170
                        exists (could_exist_alpha_subtype alpha_T) Ts)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   171
  | could_exist_alpha_subtype alpha_T T = (T = alpha_T)
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   172
fun could_exist_alpha_sub_mtype _ (alpha_T as TFree _) T =
34121
5e831d805118 get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents: 33982
diff changeset
   173
    could_exist_alpha_subtype alpha_T T
37256
0dca1ec52999 thread along context instead of theory for typedef lookup
blanchet
parents: 36385
diff changeset
   174
  | could_exist_alpha_sub_mtype ctxt alpha_T T =
0dca1ec52999 thread along context instead of theory for typedef lookup
blanchet
parents: 36385
diff changeset
   175
    (T = alpha_T orelse is_datatype ctxt [(NONE, true)] T)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   176
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   177
fun exists_alpha_sub_mtype MAlpha = true
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   178
  | exists_alpha_sub_mtype (MFun (M1, _, M2)) =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   179
    exists exists_alpha_sub_mtype [M1, M2]
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   180
  | exists_alpha_sub_mtype (MPair (M1, M2)) =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   181
    exists exists_alpha_sub_mtype [M1, M2]
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   182
  | exists_alpha_sub_mtype (MType (_, Ms)) = exists exists_alpha_sub_mtype Ms
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   183
  | exists_alpha_sub_mtype (MRec _) = true
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   184
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   185
fun exists_alpha_sub_mtype_fresh MAlpha = true
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   186
  | exists_alpha_sub_mtype_fresh (MFun (_, V _, _)) = true
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   187
  | exists_alpha_sub_mtype_fresh (MFun (_, _, M2)) =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   188
    exists_alpha_sub_mtype_fresh M2
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   189
  | exists_alpha_sub_mtype_fresh (MPair (M1, M2)) =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   190
    exists exists_alpha_sub_mtype_fresh [M1, M2]
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   191
  | exists_alpha_sub_mtype_fresh (MType (_, Ms)) =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   192
    exists exists_alpha_sub_mtype_fresh Ms
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   193
  | exists_alpha_sub_mtype_fresh (MRec _) = true
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   194
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   195
fun constr_mtype_for_binders z Ms =
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   196
  fold_rev (fn M => curry3 MFun M (A Gen)) Ms (MRec z)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   197
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   198
fun repair_mtype _ _ MAlpha = MAlpha
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   199
  | repair_mtype cache seen (MFun (M1, a, M2)) =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   200
    MFun (repair_mtype cache seen M1, a, repair_mtype cache seen M2)
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   201
  | repair_mtype cache seen (MPair Mp) =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   202
    MPair (pairself (repair_mtype cache seen) Mp)
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   203
  | repair_mtype cache seen (MType (s, Ms)) =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   204
    MType (s, maps (flatten_mtype o repair_mtype cache seen) Ms)
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   205
  | repair_mtype cache seen (MRec (z as (s, _))) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   206
    case AList.lookup (op =) cache z |> the of
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   207
      MRec _ => MType (s, [])
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   208
    | M => if member (op =) seen M then MType (s, [])
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   209
           else repair_mtype cache (M :: seen) M
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   210
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   211
fun repair_datatype_mcache cache =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   212
  let
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   213
    fun repair_one (z, M) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   214
      Unsynchronized.change cache
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   215
          (AList.update (op =) (z, repair_mtype (!cache) [] M))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   216
  in List.app repair_one (rev (!cache)) end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   217
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   218
fun repair_constr_mcache dtype_cache constr_mcache =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   219
  let
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   220
    fun repair_one (x, M) =
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   221
      Unsynchronized.change constr_mcache
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   222
          (AList.update (op =) (x, repair_mtype dtype_cache [] M))
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   223
  in List.app repair_one (!constr_mcache) end
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   224
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   225
fun is_fin_fun_supported_type @{typ prop} = true
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   226
  | is_fin_fun_supported_type @{typ bool} = true
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   227
  | is_fin_fun_supported_type (Type (@{type_name option}, _)) = true
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   228
  | is_fin_fun_supported_type _ = false
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   229
fun fin_fun_body _ _ (t as @{term False}) = SOME t
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   230
  | fin_fun_body _ _ (t as Const (@{const_name None}, _)) = SOME t
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   231
  | fin_fun_body dom_T ran_T
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   232
                 ((t0 as Const (@{const_name If}, _))
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38795
diff changeset
   233
                  $ (t1 as Const (@{const_name HOL.eq}, _) $ Bound 0 $ t1')
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   234
                  $ t2 $ t3) =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   235
    (if loose_bvar1 (t1', 0) then
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   236
       NONE
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   237
     else case fin_fun_body dom_T ran_T t3 of
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   238
       NONE => NONE
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   239
     | SOME t3 =>
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   240
       SOME (t0 $ (Const (@{const_name is_unknown}, dom_T --> bool_T) $ t1')
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   241
                $ (Const (@{const_name unknown}, ran_T)) $ (t0 $ t1 $ t2 $ t3)))
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   242
  | fin_fun_body _ _ _ = NONE
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   243
35672
ff484d4f2e14 more work on Nitpick's finite sets
blanchet
parents: 35671
diff changeset
   244
fun fresh_mfun_for_fun_type (mdata as {max_fresh, ...} : mdata) all_minus
ff484d4f2e14 more work on Nitpick's finite sets
blanchet
parents: 35671
diff changeset
   245
                            T1 T2 =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   246
  let
35672
ff484d4f2e14 more work on Nitpick's finite sets
blanchet
parents: 35671
diff changeset
   247
    val M1 = fresh_mtype_for_type mdata all_minus T1
ff484d4f2e14 more work on Nitpick's finite sets
blanchet
parents: 35671
diff changeset
   248
    val M2 = fresh_mtype_for_type mdata all_minus T2
ff484d4f2e14 more work on Nitpick's finite sets
blanchet
parents: 35671
diff changeset
   249
    val a = if not all_minus andalso exists_alpha_sub_mtype_fresh M1 andalso
ff484d4f2e14 more work on Nitpick's finite sets
blanchet
parents: 35671
diff changeset
   250
               is_fin_fun_supported_type (body_type T2) then
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   251
              V (Unsynchronized.inc max_fresh)
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   252
            else
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   253
              A Gen
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   254
  in (M1, a, M2) end
37256
0dca1ec52999 thread along context instead of theory for typedef lookup
blanchet
parents: 36385
diff changeset
   255
and fresh_mtype_for_type (mdata as {hol_ctxt as {ctxt, ...}, binarize, alpha_T,
35672
ff484d4f2e14 more work on Nitpick's finite sets
blanchet
parents: 35671
diff changeset
   256
                                    datatype_mcache, constr_mcache, ...})
ff484d4f2e14 more work on Nitpick's finite sets
blanchet
parents: 35671
diff changeset
   257
                         all_minus =
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   258
  let
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   259
    fun do_type T =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   260
      if T = alpha_T then
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   261
        MAlpha
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   262
      else case T of
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   263
        Type (@{type_name fun}, [T1, T2]) =>
39342
1a0e6f16a91b correctly thread parameter through
blanchet
parents: 39316
diff changeset
   264
        MFun (fresh_mfun_for_fun_type mdata all_minus T1 T2)
38190
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 38179
diff changeset
   265
      | Type (@{type_name prod}, [T1, T2]) => MPair (pairself do_type (T1, T2))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   266
      | Type (z as (s, _)) =>
37256
0dca1ec52999 thread along context instead of theory for typedef lookup
blanchet
parents: 36385
diff changeset
   267
        if could_exist_alpha_sub_mtype ctxt alpha_T T then
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   268
          case AList.lookup (op =) (!datatype_mcache) z of
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   269
            SOME M => M
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   270
          | NONE =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   271
            let
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   272
              val _ = Unsynchronized.change datatype_mcache (cons (z, MRec z))
35219
15a5f213ef5b fix bug in Nitpick's monotonicity code w.r.t. binary integers
blanchet
parents: 35190
diff changeset
   273
              val xs = binarized_and_boxed_datatype_constrs hol_ctxt binarize T
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   274
              val (all_Ms, constr_Ms) =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   275
                fold_rev (fn (_, T') => fn (all_Ms, constr_Ms) =>
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   276
                             let
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   277
                               val binder_Ms = map do_type (binder_types T')
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   278
                               val new_Ms = filter exists_alpha_sub_mtype_fresh
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   279
                                                   binder_Ms
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   280
                               val constr_M = constr_mtype_for_binders z
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   281
                                                                       binder_Ms
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   282
                             in
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   283
                               (union (op =) new_Ms all_Ms,
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   284
                                constr_M :: constr_Ms)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   285
                             end)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   286
                         xs ([], [])
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   287
              val M = MType (s, all_Ms)
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   288
              val _ = Unsynchronized.change datatype_mcache
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   289
                          (AList.update (op =) (z, M))
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   290
              val _ = Unsynchronized.change constr_mcache
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   291
                          (append (xs ~~ constr_Ms))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   292
            in
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   293
              if forall (not o is_MRec o snd) (!datatype_mcache) then
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   294
                (repair_datatype_mcache datatype_mcache;
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   295
                 repair_constr_mcache (!datatype_mcache) constr_mcache;
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   296
                 AList.lookup (op =) (!datatype_mcache) z |> the)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   297
              else
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   298
                M
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   299
            end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   300
        else
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   301
          MType (s, [])
37260
dde817e6dfb1 added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents: 37256
diff changeset
   302
      | _ => MType (simple_string_of_typ T, [])
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   303
  in do_type end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   304
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   305
fun prodM_factors (MPair (M1, M2)) = maps prodM_factors [M1, M2]
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   306
  | prodM_factors M = [M]
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   307
fun curried_strip_mtype (MFun (M1, _, M2)) =
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   308
    curried_strip_mtype M2 |>> append (prodM_factors M1)
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   309
  | curried_strip_mtype M = ([], M)
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   310
fun sel_mtype_from_constr_mtype s M =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   311
  let val (arg_Ms, dataM) = curried_strip_mtype M in
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   312
    MFun (dataM, A Gen,
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   313
          case sel_no_from_name s of ~1 => bool_M | n => nth arg_Ms n)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   314
  end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   315
37256
0dca1ec52999 thread along context instead of theory for typedef lookup
blanchet
parents: 36385
diff changeset
   316
fun mtype_for_constr (mdata as {hol_ctxt = {ctxt, ...}, alpha_T, constr_mcache,
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   317
                                ...}) (x as (_, T)) =
37256
0dca1ec52999 thread along context instead of theory for typedef lookup
blanchet
parents: 36385
diff changeset
   318
  if could_exist_alpha_sub_mtype ctxt alpha_T T then
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   319
    case AList.lookup (op =) (!constr_mcache) x of
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   320
      SOME M => M
35384
88dbcfe75c45 cosmetics
blanchet
parents: 35280
diff changeset
   321
    | NONE => if T = alpha_T then
35672
ff484d4f2e14 more work on Nitpick's finite sets
blanchet
parents: 35671
diff changeset
   322
                let val M = fresh_mtype_for_type mdata false T in
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   323
                  (Unsynchronized.change constr_mcache (cons (x, M)); M)
35384
88dbcfe75c45 cosmetics
blanchet
parents: 35280
diff changeset
   324
                end
88dbcfe75c45 cosmetics
blanchet
parents: 35280
diff changeset
   325
              else
35672
ff484d4f2e14 more work on Nitpick's finite sets
blanchet
parents: 35671
diff changeset
   326
                (fresh_mtype_for_type mdata false (body_type T);
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   327
                 AList.lookup (op =) (!constr_mcache) x |> the)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   328
  else
35672
ff484d4f2e14 more work on Nitpick's finite sets
blanchet
parents: 35671
diff changeset
   329
    fresh_mtype_for_type mdata false T
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   330
fun mtype_for_sel (mdata as {hol_ctxt, binarize, ...}) (x as (s, _)) =
35190
ce653cc27a94 make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents: 35079
diff changeset
   331
  x |> binarized_and_boxed_constr_for_sel hol_ctxt binarize
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   332
    |> mtype_for_constr mdata |> sel_mtype_from_constr_mtype s
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   333
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   334
fun resolve_annotation_atom lits (V x) =
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   335
    x |> AList.lookup (op =) lits |> Option.map A |> the_default (V x)
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   336
  | resolve_annotation_atom _ a = a
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   337
fun resolve_mtype lits =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   338
  let
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   339
    fun aux MAlpha = MAlpha
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   340
      | aux (MFun (M1, a, M2)) =
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   341
        MFun (aux M1, resolve_annotation_atom lits a, aux M2)
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   342
      | aux (MPair Mp) = MPair (pairself aux Mp)
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   343
      | aux (MType (s, Ms)) = MType (s, map aux Ms)
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   344
      | aux (MRec z) = MRec z
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   345
  in aux end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   346
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   347
datatype comp_op = Eq | Leq
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   348
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   349
type comp = annotation_atom * annotation_atom * comp_op * var list
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   350
type annotation_expr = literal list
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   351
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   352
type constraint_set = literal list * comp list * annotation_expr list
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   353
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   354
fun string_for_comp_op Eq = "="
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   355
  | string_for_comp_op Leq = "\<le>"
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   356
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   357
fun string_for_annotation_expr [] = "\<bot>"
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   358
  | string_for_annotation_expr lits =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   359
    space_implode " \<or> " (map string_for_literal lits)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   360
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   361
fun do_literal _ NONE = NONE
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   362
  | do_literal (x, a) (SOME lits) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   363
    case AList.lookup (op =) lits x of
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   364
      SOME a' => if a = a' then SOME lits else NONE
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   365
    | NONE => SOME ((x, a) :: lits)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   366
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   367
fun do_annotation_atom_comp Eq [] a1 a2 (accum as (lits, comps)) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   368
    (case (a1, a2) of
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   369
       (A sn1, A sn2) => if sn1 = sn2 then SOME accum else NONE
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   370
     | (V x1, A sn2) =>
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   371
       Option.map (rpair comps) (do_literal (x1, sn2) (SOME lits))
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   372
     | (V _, V _) => SOME (lits, insert (op =) (a1, a2, Eq, []) comps)
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   373
     | _ => do_annotation_atom_comp Eq [] a2 a1 accum)
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   374
  | do_annotation_atom_comp Leq [] a1 a2 (accum as (lits, comps)) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   375
    (case (a1, a2) of
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   376
       (_, A Gen) => SOME accum
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   377
     | (A Fls, _) => SOME accum
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   378
     | (A Gen, A Fls) => NONE
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   379
     | (V _, V _) => SOME (lits, insert (op =) (a1, a2, Leq, []) comps)
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   380
     | _ => do_annotation_atom_comp Eq [] a1 a2 accum)
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   381
  | do_annotation_atom_comp cmp xs a1 a2 (lits, comps) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   382
    SOME (lits, insert (op =) (a1, a2, cmp, xs) comps)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   383
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   384
fun do_mtype_comp _ _ _ _ NONE = NONE
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   385
  | do_mtype_comp _ _ MAlpha MAlpha accum = accum
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   386
  | do_mtype_comp Eq xs (MFun (M11, a1, M12)) (MFun (M21, a2, M22))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   387
                  (SOME accum) =
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   388
     accum |> do_annotation_atom_comp Eq xs a1 a2 |> do_mtype_comp Eq xs M11 M21
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   389
           |> do_mtype_comp Eq xs M12 M22
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   390
  | do_mtype_comp Leq xs (MFun (M11, a1, M12)) (MFun (M21, a2, M22))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   391
                  (SOME accum) =
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   392
    (if exists_alpha_sub_mtype M11 then
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   393
       accum |> do_annotation_atom_comp Leq xs a1 a2
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   394
             |> do_mtype_comp Leq xs M21 M11
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   395
             |> (case a2 of
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   396
                   A Gen => I
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   397
                 | A Fls => do_mtype_comp Leq xs M11 M21
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   398
                 | V x => do_mtype_comp Leq (x :: xs) M11 M21)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   399
     else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   400
       SOME accum)
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   401
    |> do_mtype_comp Leq xs M12 M22
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   402
  | do_mtype_comp cmp xs (M1 as MPair (M11, M12)) (M2 as MPair (M21, M22))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   403
                  accum =
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   404
    (accum |> fold (uncurry (do_mtype_comp cmp xs)) [(M11, M21), (M12, M22)]
40722
441260986b63 make two copies (!) of Library.UnequalLengths coincide with ListPair.UnequalLengths;
wenzelm
parents: 40146
diff changeset
   405
     handle ListPair.UnequalLengths =>
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   406
            raise MTYPE ("Nitpick_Mono.do_mtype_comp", [M1, M2], []))
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   407
  | do_mtype_comp _ _ (MType _) (MType _) accum =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   408
    accum (* no need to compare them thanks to the cache *)
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   409
  | do_mtype_comp cmp _ M1 M2 _ =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   410
    raise MTYPE ("Nitpick_Mono.do_mtype_comp (" ^ string_for_comp_op cmp ^ ")",
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   411
                 [M1, M2], [])
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   412
35832
1dac16f00cd2 added type constraints to make SML/NJ happy
blanchet
parents: 35812
diff changeset
   413
fun add_mtype_comp cmp M1 M2 ((lits, comps, sexps) : constraint_set) =
38179
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   414
    (trace_msg (fn () => "*** Add " ^ string_for_mtype M1 ^ " " ^
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   415
                         string_for_comp_op cmp ^ " " ^ string_for_mtype M2);
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   416
     case do_mtype_comp cmp [] M1 M2 (SOME (lits, comps)) of
38179
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   417
       NONE => (trace_msg (K "**** Unsolvable"); raise UNSOLVABLE ())
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
   418
     | SOME (lits, comps) => (lits, comps, sexps))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   419
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   420
val add_mtypes_equal = add_mtype_comp Eq
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   421
val add_is_sub_mtype = add_mtype_comp Leq
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   422
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   423
fun do_notin_mtype_fv _ _ _ NONE = NONE
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   424
  | do_notin_mtype_fv Minus _ MAlpha accum = accum
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   425
  | do_notin_mtype_fv Plus [] MAlpha _ = NONE
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   426
  | do_notin_mtype_fv Plus [(x, sn)] MAlpha (SOME (lits, sexps)) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   427
    SOME lits |> do_literal (x, sn) |> Option.map (rpair sexps)
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   428
  | do_notin_mtype_fv Plus sexp MAlpha (SOME (lits, sexps)) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   429
    SOME (lits, insert (op =) sexp sexps)
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   430
  | do_notin_mtype_fv sn sexp (MFun (M1, A a, M2)) accum =
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   431
    accum |> (if a = Fls andalso sn = Plus then do_notin_mtype_fv Plus sexp M1
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   432
              else I)
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   433
          |> (if a = Gen orelse sn = Plus then do_notin_mtype_fv Minus sexp M1
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   434
              else I)
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   435
          |> do_notin_mtype_fv sn sexp M2
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   436
  | do_notin_mtype_fv Plus sexp (MFun (M1, V x, M2)) accum =
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   437
    accum |> (case do_literal (x, Gen) (SOME sexp) of
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   438
                NONE => I
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   439
              | SOME sexp' => do_notin_mtype_fv Plus sexp' M1)
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   440
          |> do_notin_mtype_fv Minus sexp M1
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   441
          |> do_notin_mtype_fv Plus sexp M2
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   442
  | do_notin_mtype_fv Minus sexp (MFun (M1, V x, M2)) accum =
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   443
    accum |> (case do_literal (x, Fls) (SOME sexp) of
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   444
                NONE => I
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   445
              | SOME sexp' => do_notin_mtype_fv Plus sexp' M1)
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   446
          |> do_notin_mtype_fv Minus sexp M2
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   447
  | do_notin_mtype_fv sn sexp (MPair (M1, M2)) accum =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   448
    accum |> fold (do_notin_mtype_fv sn sexp) [M1, M2]
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   449
  | do_notin_mtype_fv sn sexp (MType (_, Ms)) accum =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   450
    accum |> fold (do_notin_mtype_fv sn sexp) Ms
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   451
  | do_notin_mtype_fv _ _ M _ =
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   452
    raise MTYPE ("Nitpick_Mono.do_notin_mtype_fv", [M], [])
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   453
35832
1dac16f00cd2 added type constraints to make SML/NJ happy
blanchet
parents: 35812
diff changeset
   454
fun add_notin_mtype_fv sn M ((lits, comps, sexps) : constraint_set) =
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   455
  (trace_msg (fn () => "*** Add " ^ string_for_mtype M ^ " is " ^
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   456
                       (case sn of Minus => "concrete" | Plus => "complete"));
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   457
   case do_notin_mtype_fv sn [] M (SOME (lits, sexps)) of
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   458
     NONE => (trace_msg (K "**** Unsolvable"); raise UNSOLVABLE ())
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   459
   | SOME (lits, sexps) => (lits, comps, sexps))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   460
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   461
val add_mtype_is_concrete = add_notin_mtype_fv Minus
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   462
val add_mtype_is_complete = add_notin_mtype_fv Plus
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   463
35384
88dbcfe75c45 cosmetics
blanchet
parents: 35280
diff changeset
   464
val bool_from_minus = true
88dbcfe75c45 cosmetics
blanchet
parents: 35280
diff changeset
   465
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   466
fun bool_from_annotation Fls = not bool_from_minus
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   467
  | bool_from_annotation Gen = bool_from_minus
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   468
fun sign_from_bool b = if b = bool_from_minus then Gen else Fls
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   469
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   470
fun prop_for_literal (x, a) =
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   471
  PropLogic.BoolVar x |> not (bool_from_annotation a) ? PropLogic.Not
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   472
fun prop_for_annotation_atom_eq (A a', a) =
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   473
    if a = a' then PropLogic.True else PropLogic.False
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   474
  | prop_for_annotation_atom_eq (V x, a) = prop_for_literal (x, a)
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   475
fun prop_for_annotation_expr xs = PropLogic.exists (map prop_for_literal xs)
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   476
fun prop_for_exists_eq xs a =
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   477
  PropLogic.exists (map (fn x => prop_for_literal (x, a)) xs)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   478
fun prop_for_comp (a1, a2, Eq, []) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   479
    PropLogic.SAnd (prop_for_comp (a1, a2, Leq, []),
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   480
                    prop_for_comp (a2, a1, Leq, []))
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   481
  | prop_for_comp (a1, a2, Leq, []) =
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   482
    PropLogic.SOr (prop_for_annotation_atom_eq (a1, Fls),
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   483
                   prop_for_annotation_atom_eq (a2, Gen))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   484
  | prop_for_comp (a1, a2, cmp, xs) =
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   485
    PropLogic.SOr (prop_for_exists_eq xs Gen, prop_for_comp (a1, a2, cmp, []))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   486
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   487
fun literals_from_assignments max_var assigns lits =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   488
  fold (fn x => fn accum =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   489
           if AList.defined (op =) lits x then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   490
             accum
34123
c4988215a691 distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents: 34121
diff changeset
   491
           else case assigns x of
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   492
             SOME b => (x, sign_from_bool b) :: accum
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   493
           | NONE => accum) (max_var downto 1) lits
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   494
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   495
fun string_for_comp (a1, a2, cmp, xs) =
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   496
  string_for_annotation_atom a1 ^ " " ^ string_for_comp_op cmp ^
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   497
  subscript_string_for_vars " \<and> " xs ^ " " ^ string_for_annotation_atom a2
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   498
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   499
fun print_problem lits comps sexps =
38179
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   500
  trace_msg (fn () => "*** Problem:\n" ^
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   501
                      cat_lines (map string_for_literal lits @
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   502
                                 map string_for_comp comps @
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   503
                                 map string_for_annotation_expr sexps))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   504
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   505
fun print_solution lits =
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   506
  let val (fs, gs) = List.partition (curry (op =) Fls o snd) lits in
38179
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   507
    trace_msg (fn () => "*** Solution:\n" ^
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   508
                        "F: " ^ commas (map (string_for_var o fst) fs) ^ "\n" ^
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   509
                        "G: " ^ commas (map (string_for_var o fst) gs))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   510
  end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   511
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
   512
fun solve max_var (lits, comps, sexps) =
40146
f2a14b6effcf whitespace tuning
blanchet
parents: 40058
diff changeset
   513
  let
f2a14b6effcf whitespace tuning
blanchet
parents: 40058
diff changeset
   514
    fun do_assigns assigns =
f2a14b6effcf whitespace tuning
blanchet
parents: 40058
diff changeset
   515
      SOME (literals_from_assignments max_var assigns lits
f2a14b6effcf whitespace tuning
blanchet
parents: 40058
diff changeset
   516
            |> tap print_solution)
f2a14b6effcf whitespace tuning
blanchet
parents: 40058
diff changeset
   517
    val _ = print_problem lits comps sexps
f2a14b6effcf whitespace tuning
blanchet
parents: 40058
diff changeset
   518
    val prop = PropLogic.all (map prop_for_literal lits @
f2a14b6effcf whitespace tuning
blanchet
parents: 40058
diff changeset
   519
                              map prop_for_comp comps @
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   520
                              map prop_for_annotation_expr sexps)
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   521
    val default_val = bool_from_annotation Gen
40146
f2a14b6effcf whitespace tuning
blanchet
parents: 40058
diff changeset
   522
  in
f2a14b6effcf whitespace tuning
blanchet
parents: 40058
diff changeset
   523
    if PropLogic.eval (K default_val) prop then
f2a14b6effcf whitespace tuning
blanchet
parents: 40058
diff changeset
   524
      do_assigns (K (SOME default_val))
f2a14b6effcf whitespace tuning
blanchet
parents: 40058
diff changeset
   525
    else
f2a14b6effcf whitespace tuning
blanchet
parents: 40058
diff changeset
   526
      let
f2a14b6effcf whitespace tuning
blanchet
parents: 40058
diff changeset
   527
        (* use the first ML solver (to avoid startup overhead) *)
f2a14b6effcf whitespace tuning
blanchet
parents: 40058
diff changeset
   528
        val solvers = !SatSolver.solvers
f2a14b6effcf whitespace tuning
blanchet
parents: 40058
diff changeset
   529
                      |> filter (member (op =) ["dptsat", "dpll"] o fst)
f2a14b6effcf whitespace tuning
blanchet
parents: 40058
diff changeset
   530
      in
f2a14b6effcf whitespace tuning
blanchet
parents: 40058
diff changeset
   531
        case snd (hd solvers) prop of
f2a14b6effcf whitespace tuning
blanchet
parents: 40058
diff changeset
   532
          SatSolver.SATISFIABLE assigns => do_assigns assigns
f2a14b6effcf whitespace tuning
blanchet
parents: 40058
diff changeset
   533
        | _ => NONE
f2a14b6effcf whitespace tuning
blanchet
parents: 40058
diff changeset
   534
      end
f2a14b6effcf whitespace tuning
blanchet
parents: 40058
diff changeset
   535
  end
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   536
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   537
type mtype_schema = mtyp * constraint_set
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   538
type mtype_context =
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   539
  {bound_Ts: typ list,
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   540
   bound_Ms: mtyp list,
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   541
   frees: (styp * mtyp) list,
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   542
   consts: (styp * mtyp) list}
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   543
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   544
type accumulator = mtype_context * constraint_set
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   545
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   546
val initial_gamma = {bound_Ts = [], bound_Ms = [], frees = [], consts = []}
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   547
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   548
fun push_bound T M {bound_Ts, bound_Ms, frees, consts} =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   549
  {bound_Ts = T :: bound_Ts, bound_Ms = M :: bound_Ms, frees = frees,
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   550
   consts = consts}
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   551
fun pop_bound {bound_Ts, bound_Ms, frees, consts} =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   552
  {bound_Ts = tl bound_Ts, bound_Ms = tl bound_Ms, frees = frees,
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   553
   consts = consts}
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   554
  handle List.Empty => initial_gamma (* FIXME: needed? *)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   555
39359
6f49c7fbb1b1 remove "fast_descs" option from Nitpick;
blanchet
parents: 39345
diff changeset
   556
fun consider_term (mdata as {hol_ctxt = {thy, ctxt, stds, ...}, alpha_T,
6f49c7fbb1b1 remove "fast_descs" option from Nitpick;
blanchet
parents: 39345
diff changeset
   557
                             max_fresh, ...}) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   558
  let
37476
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
   559
    fun is_enough_eta_expanded t =
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
   560
      case strip_comb t of
39359
6f49c7fbb1b1 remove "fast_descs" option from Nitpick;
blanchet
parents: 39345
diff changeset
   561
        (Const x, ts) => the_default 0 (arity_of_built_in_const thy stds x)
37476
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
   562
        <= length ts
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
   563
      | _ => true
35672
ff484d4f2e14 more work on Nitpick's finite sets
blanchet
parents: 35671
diff changeset
   564
    val mtype_for = fresh_mtype_for_type mdata false
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   565
    fun plus_set_mtype_for_dom M =
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   566
      MFun (M, A (if exists_alpha_sub_mtype M then Fls else Gen), bool_M)
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   567
    fun do_all T (gamma, cset) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   568
      let
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   569
        val abs_M = mtype_for (domain_type (domain_type T))
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   570
        val body_M = mtype_for (body_type T)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   571
      in
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   572
        (MFun (MFun (abs_M, A Gen, body_M), A Gen, body_M),
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   573
         (gamma, cset |> add_mtype_is_complete abs_M))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   574
      end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   575
    fun do_equals T (gamma, cset) =
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   576
      let val M = mtype_for (domain_type T) in
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   577
        (MFun (M, A Gen, MFun (M, V (Unsynchronized.inc max_fresh),
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   578
                               mtype_for (nth_range_type 2 T))),
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   579
         (gamma, cset |> add_mtype_is_concrete M))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   580
      end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   581
    fun do_robust_set_operation T (gamma, cset) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   582
      let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   583
        val set_T = domain_type T
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   584
        val M1 = mtype_for set_T
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   585
        val M2 = mtype_for set_T
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   586
        val M3 = mtype_for set_T
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   587
      in
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   588
        (MFun (M1, A Gen, MFun (M2, A Gen, M3)),
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   589
         (gamma, cset |> add_is_sub_mtype M1 M3 |> add_is_sub_mtype M2 M3))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   590
      end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   591
    fun do_fragile_set_operation T (gamma, cset) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   592
      let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   593
        val set_T = domain_type T
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   594
        val set_M = mtype_for set_T
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   595
        fun custom_mtype_for (T as Type (@{type_name fun}, [T1, T2])) =
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   596
            if T = set_T then set_M
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   597
            else MFun (custom_mtype_for T1, A Gen, custom_mtype_for T2)
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   598
          | custom_mtype_for T = mtype_for T
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   599
      in
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   600
        (custom_mtype_for T, (gamma, cset |> add_mtype_is_concrete set_M))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   601
      end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   602
    fun do_pair_constr T accum =
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   603
      case mtype_for (nth_range_type 2 T) of
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   604
        M as MPair (a_M, b_M) =>
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   605
        (MFun (a_M, A Gen, MFun (b_M, A Gen, M)), accum)
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   606
      | M => raise MTYPE ("Nitpick_Mono.consider_term.do_pair_constr", [M], [])
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   607
    fun do_nth_pair_sel n T =
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   608
      case mtype_for (domain_type T) of
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   609
        M as MPair (a_M, b_M) =>
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   610
        pair (MFun (M, A Gen, if n = 0 then a_M else b_M))
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   611
      | M => raise MTYPE ("Nitpick_Mono.consider_term.do_nth_pair_sel", [M], [])
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   612
    fun do_bounded_quantifier t0 abs_s abs_T connective_t bound_t body_t accum =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   613
      let
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   614
        val abs_M = mtype_for abs_T
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   615
        val (bound_m, accum) =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   616
          accum |>> push_bound abs_T abs_M |> do_term bound_t
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   617
        val expected_bound_M = plus_set_mtype_for_dom abs_M
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   618
        val (body_m, accum) =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   619
          accum ||> add_mtypes_equal expected_bound_M (mtype_of_mterm bound_m)
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   620
                |> do_term body_t ||> apfst pop_bound
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   621
        val bound_M = mtype_of_mterm bound_m
39345
062c10ff848c remove unreferenced identifiers
blanchet
parents: 39342
diff changeset
   622
        val (M1, a, _) = dest_MFun bound_M
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   623
      in
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   624
        (MApp (MRaw (t0, MFun (bound_M, A Gen, bool_M)),
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   625
               MAbs (abs_s, abs_T, M1, a,
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   626
                     MApp (MApp (MRaw (connective_t,
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   627
                                       mtype_for (fastype_of connective_t)),
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   628
                                 MApp (bound_m, MRaw (Bound 0, M1))),
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   629
                           body_m))), accum)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   630
      end
39345
062c10ff848c remove unreferenced identifiers
blanchet
parents: 39342
diff changeset
   631
    and do_term t (accum as ({bound_Ts, bound_Ms, frees, consts}, cset)) =
38179
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   632
        (trace_msg (fn () => "  \<Gamma> \<turnstile> " ^
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   633
                             Syntax.string_of_term ctxt t ^ " : _?");
37704
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   634
         case t of
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   635
           Const (x as (s, T)) =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   636
           (case AList.lookup (op =) consts x of
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   637
              SOME M => (M, accum)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   638
            | NONE =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   639
              if not (could_exist_alpha_subtype alpha_T T) then
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   640
                (mtype_for T, accum)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   641
              else case s of
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   642
                @{const_name all} => do_all T accum
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   643
              | @{const_name "=="} => do_equals T accum
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   644
              | @{const_name All} => do_all T accum
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   645
              | @{const_name Ex} =>
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   646
                let val set_T = domain_type T in
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   647
                  do_term (Abs (Name.uu, set_T,
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   648
                                @{const Not} $ (HOLogic.mk_eq
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   649
                                    (Abs (Name.uu, domain_type set_T,
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   650
                                          @{const False}),
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   651
                                     Bound 0)))) accum
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   652
                  |>> mtype_of_mterm
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   653
                end
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38795
diff changeset
   654
              | @{const_name HOL.eq} => do_equals T accum
37476
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
   655
              | @{const_name The} =>
38179
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   656
                (trace_msg (K "*** The"); raise UNSOLVABLE ())
37476
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
   657
              | @{const_name Eps} =>
38179
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   658
                (trace_msg (K "*** Eps"); raise UNSOLVABLE ())
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   659
              | @{const_name If} =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   660
                do_robust_set_operation (range_type T) accum
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   661
                |>> curry3 MFun bool_M (A Gen)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   662
              | @{const_name Pair} => do_pair_constr T accum
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   663
              | @{const_name fst} => do_nth_pair_sel 0 T accum
39316
b6c4385ab400 change defaults of Auto Nitpick so that it consumes less resources (time and Kodkod threads)
blanchet
parents: 38864
diff changeset
   664
              | @{const_name snd} => do_nth_pair_sel 1 T accum
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   665
              | @{const_name Id} =>
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   666
                (MFun (mtype_for (domain_type T), A Gen, bool_M), accum)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   667
              | @{const_name converse} =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   668
                let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   669
                  val x = Unsynchronized.inc max_fresh
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   670
                  fun mtype_for_set T =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   671
                    MFun (mtype_for (domain_type T), V x, bool_M)
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   672
                  val ab_set_M = domain_type T |> mtype_for_set
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   673
                  val ba_set_M = range_type T |> mtype_for_set
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   674
                in (MFun (ab_set_M, A Gen, ba_set_M), accum) end
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   675
              | @{const_name trancl} => do_fragile_set_operation T accum
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   676
              | @{const_name rel_comp} =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   677
                let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   678
                  val x = Unsynchronized.inc max_fresh
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   679
                  fun mtype_for_set T =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   680
                    MFun (mtype_for (domain_type T), V x, bool_M)
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   681
                  val bc_set_M = domain_type T |> mtype_for_set
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   682
                  val ab_set_M = domain_type (range_type T) |> mtype_for_set
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   683
                  val ac_set_M = nth_range_type 2 T |> mtype_for_set
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   684
                in
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   685
                  (MFun (bc_set_M, A Gen, MFun (ab_set_M, A Gen, ac_set_M)),
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   686
                   accum)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   687
                end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   688
              | @{const_name image} =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   689
                let
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   690
                  val a_M = mtype_for (domain_type (domain_type T))
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   691
                  val b_M = mtype_for (range_type (domain_type T))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   692
                in
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   693
                  (MFun (MFun (a_M, A Gen, b_M), A Gen,
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   694
                         MFun (plus_set_mtype_for_dom a_M, A Gen,
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   695
                               plus_set_mtype_for_dom b_M)), accum)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   696
                end
35672
ff484d4f2e14 more work on Nitpick's finite sets
blanchet
parents: 35671
diff changeset
   697
              | @{const_name finite} =>
ff484d4f2e14 more work on Nitpick's finite sets
blanchet
parents: 35671
diff changeset
   698
                let val M1 = mtype_for (domain_type (domain_type T)) in
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   699
                  (MFun (plus_set_mtype_for_dom M1, A Gen, bool_M), accum)
35672
ff484d4f2e14 more work on Nitpick's finite sets
blanchet
parents: 35671
diff changeset
   700
                end
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   701
              | @{const_name Sigma} =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   702
                let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   703
                  val x = Unsynchronized.inc max_fresh
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   704
                  fun mtype_for_set T =
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   705
                    MFun (mtype_for (domain_type T), V x, bool_M)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   706
                  val a_set_T = domain_type T
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   707
                  val a_M = mtype_for (domain_type a_set_T)
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   708
                  val b_set_M = mtype_for_set (range_type (domain_type
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   709
                                                               (range_type T)))
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   710
                  val a_set_M = mtype_for_set a_set_T
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   711
                  val a_to_b_set_M = MFun (a_M, A Gen, b_set_M)
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   712
                  val ab_set_M = mtype_for_set (nth_range_type 2 T)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   713
                in
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   714
                  (MFun (a_set_M, A Gen, MFun (a_to_b_set_M, A Gen, ab_set_M)),
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   715
                   accum)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   716
                end
35220
2bcdae5f4fdb added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents: 35219
diff changeset
   717
              | _ =>
39360
cdf2c3341422 eliminate more clutter related to "fast_descrs" optimization
blanchet
parents: 39359
diff changeset
   718
                if s = @{const_name safe_The} then
35671
ed2c3830d881 improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents: 35665
diff changeset
   719
                  let
ed2c3830d881 improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents: 35665
diff changeset
   720
                    val a_set_M = mtype_for (domain_type T)
ed2c3830d881 improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents: 35665
diff changeset
   721
                    val a_M = dest_MFun a_set_M |> #1
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   722
                  in (MFun (a_set_M, A Gen, a_M), accum) end
35220
2bcdae5f4fdb added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents: 35219
diff changeset
   723
                else if s = @{const_name ord_class.less_eq} andalso
2bcdae5f4fdb added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents: 35219
diff changeset
   724
                        is_set_type (domain_type T) then
2bcdae5f4fdb added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents: 35219
diff changeset
   725
                  do_fragile_set_operation T accum
2bcdae5f4fdb added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents: 35219
diff changeset
   726
                else if is_sel s then
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   727
                  (mtype_for_sel mdata x, accum)
37256
0dca1ec52999 thread along context instead of theory for typedef lookup
blanchet
parents: 36385
diff changeset
   728
                else if is_constr ctxt stds x then
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   729
                  (mtype_for_constr mdata x, accum)
39359
6f49c7fbb1b1 remove "fast_descs" option from Nitpick;
blanchet
parents: 39345
diff changeset
   730
                else if is_built_in_const thy stds x then
35672
ff484d4f2e14 more work on Nitpick's finite sets
blanchet
parents: 35671
diff changeset
   731
                  (fresh_mtype_for_type mdata true T, accum)
35220
2bcdae5f4fdb added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents: 35219
diff changeset
   732
                else
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   733
                  let val M = mtype_for T in
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   734
                    (M, ({bound_Ts = bound_Ts, bound_Ms = bound_Ms,
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   735
                          frees = frees, consts = (x, M) :: consts}, cset))
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   736
                  end) |>> curry MRaw t
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   737
         | Free (x as (_, T)) =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   738
           (case AList.lookup (op =) frees x of
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   739
              SOME M => (M, accum)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   740
            | NONE =>
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   741
              let val M = mtype_for T in
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   742
                (M, ({bound_Ts = bound_Ts, bound_Ms = bound_Ms,
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   743
                      frees = (x, M) :: frees, consts = consts}, cset))
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   744
              end) |>> curry MRaw t
38179
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   745
         | Var _ => (trace_msg (K "*** Var"); raise UNSOLVABLE ())
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   746
         | Bound j => (MRaw (t, nth bound_Ms j), accum)
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   747
         | Abs (s, T, t') =>
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   748
           (case fin_fun_body T (fastype_of1 (T :: bound_Ts, t')) t' of
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   749
              SOME t' =>
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   750
              let
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   751
                val M = mtype_for T
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   752
                val a = V (Unsynchronized.inc max_fresh)
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   753
                val (m', accum) = do_term t' (accum |>> push_bound T M)
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   754
              in (MAbs (s, T, M, a, m'), accum |>> pop_bound) end
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   755
            | NONE =>
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   756
              ((case t' of
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   757
                  t1' $ Bound 0 =>
37476
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
   758
                  if not (loose_bvar1 (t1', 0)) andalso
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
   759
                     is_enough_eta_expanded t1' then
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   760
                    do_term (incr_boundvars ~1 t1') accum
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   761
                  else
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   762
                    raise SAME ()
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38795
diff changeset
   763
                | (t11 as Const (@{const_name HOL.eq}, _)) $ Bound 0 $ t13 =>
37476
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
   764
                  if not (loose_bvar1 (t13, 0)) then
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
   765
                    do_term (incr_boundvars ~1 (t11 $ t13)) accum
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
   766
                  else
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
   767
                    raise SAME ()
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   768
                | _ => raise SAME ())
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   769
               handle SAME () =>
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   770
                      let
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   771
                        val M = mtype_for T
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   772
                        val (m', accum) = do_term t' (accum |>> push_bound T M)
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   773
                      in (MAbs (s, T, M, A Gen, m'), accum |>> pop_bound) end))
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   774
         | (t0 as Const (@{const_name All}, _))
38786
e46e7a9cb622 formerly unnamed infix impliciation now named HOL.implies
haftmann
parents: 38199
diff changeset
   775
           $ Abs (s', T', (t10 as @{const HOL.implies}) $ (t11 $ Bound 0) $ t12) =>
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   776
           do_bounded_quantifier t0 s' T' t10 t11 t12 accum
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   777
         | (t0 as Const (@{const_name Ex}, _))
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   778
           $ Abs (s', T', (t10 as @{const HOL.conj}) $ (t11 $ Bound 0) $ t12) =>
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   779
           do_bounded_quantifier t0 s' T' t10 t11 t12 accum
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   780
         | Const (@{const_name Let}, _) $ t1 $ t2 =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   781
           do_term (betapply (t2, t1)) accum
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   782
         | t1 $ t2 =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   783
           let
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   784
             val (m1, accum) = do_term t1 accum
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   785
             val (m2, accum) = do_term t2 accum
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   786
           in
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
   787
             let
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
   788
               val M11 = mtype_of_mterm m1 |> dest_MFun |> #1
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
   789
               val M2 = mtype_of_mterm m2
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
   790
             in (MApp (m1, m2), accum ||> add_is_sub_mtype M2 M11) end
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   791
           end)
38179
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   792
        |> tap (fn (m, _) => trace_msg (fn () => "  \<Gamma> \<turnstile> " ^
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   793
                                                 string_for_mterm ctxt m))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   794
  in do_term end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   795
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   796
fun force_minus_funs 0 _ = I
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   797
  | force_minus_funs n (M as MFun (M1, _, M2)) =
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   798
    add_mtypes_equal M (MFun (M1, A Gen, M2)) #> force_minus_funs (n - 1) M2
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   799
  | force_minus_funs _ M =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   800
    raise MTYPE ("Nitpick_Mono.force_minus_funs", [M], [])
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   801
fun consider_general_equals mdata def (x as (_, T)) t1 t2 accum =
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   802
  let
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   803
    val (m1, accum) = consider_term mdata t1 accum
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   804
    val (m2, accum) = consider_term mdata t2 accum
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   805
    val M1 = mtype_of_mterm m1
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   806
    val M2 = mtype_of_mterm m2
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   807
    val accum = accum ||> add_mtypes_equal M1 M2
35672
ff484d4f2e14 more work on Nitpick's finite sets
blanchet
parents: 35671
diff changeset
   808
    val body_M = fresh_mtype_for_type mdata false (nth_range_type 2 T)
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   809
    val m = MApp (MApp (MRaw (Const x,
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   810
                           MFun (M1, A Gen, MFun (M2, A Gen, body_M))), m1), m2)
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   811
  in
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   812
    (m, if def then
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   813
          let val (head_m, arg_ms) = strip_mcomb m1 in
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   814
            accum ||> force_minus_funs (length arg_ms) (mtype_of_mterm head_m)
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   815
          end
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   816
        else
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   817
          accum)
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   818
  end
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   819
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   820
fun consider_general_formula (mdata as {hol_ctxt = {ctxt, ...}, ...}) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   821
  let
35672
ff484d4f2e14 more work on Nitpick's finite sets
blanchet
parents: 35671
diff changeset
   822
    val mtype_for = fresh_mtype_for_type mdata false
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   823
    val do_term = consider_term mdata
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
   824
    fun do_formula sn t accum =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   825
        let
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   826
          fun do_quantifier (quant_x as (quant_s, _)) abs_s abs_T body_t =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   827
            let
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   828
              val abs_M = mtype_for abs_T
34982
7b8c366e34a2 added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents: 34936
diff changeset
   829
              val side_cond = ((sn = Minus) = (quant_s = @{const_name Ex}))
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   830
              val (body_m, accum) =
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   831
                accum ||> side_cond ? add_mtype_is_complete abs_M
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   832
                      |>> push_bound abs_T abs_M |> do_formula sn body_t
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   833
              val body_M = mtype_of_mterm body_m
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   834
            in
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   835
              (MApp (MRaw (Const quant_x,
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   836
                           MFun (MFun (abs_M, A Gen, body_M), A Gen, body_M)),
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   837
                     MAbs (abs_s, abs_T, abs_M, A Gen, body_m)),
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   838
               accum |>> pop_bound)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   839
            end
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   840
          fun do_equals x t1 t2 =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   841
            case sn of
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   842
              Plus => do_term t accum
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   843
            | Minus => consider_general_equals mdata false x t1 t2 accum
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   844
        in
38179
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   845
          (trace_msg (fn () => "  \<Gamma> \<turnstile> " ^
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   846
                               Syntax.string_of_term ctxt t ^ " : o\<^sup>" ^
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   847
                               string_for_sign sn ^ "?");
37704
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   848
           case t of
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   849
             Const (x as (@{const_name all}, _)) $ Abs (s1, T1, t1) =>
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   850
             do_quantifier x s1 T1 t1
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   851
           | Const (x as (@{const_name "=="}, _)) $ t1 $ t2 => do_equals x t1 t2
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   852
           | @{const Trueprop} $ t1 =>
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   853
             let val (m1, accum) = do_formula sn t1 accum in
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   854
               (MApp (MRaw (@{const Trueprop}, mtype_for (bool_T --> prop_T)),
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   855
                      m1), accum)
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   856
             end
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   857
           | @{const Not} $ t1 =>
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   858
             let val (m1, accum) = do_formula (negate_sign sn) t1 accum in
37704
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   859
               (MApp (MRaw (@{const Not}, mtype_for (bool_T --> bool_T)), m1),
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   860
                accum)
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   861
             end
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   862
           | Const (x as (@{const_name All}, _)) $ Abs (s1, T1, t1) =>
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   863
             do_quantifier x s1 T1 t1
39345
062c10ff848c remove unreferenced identifiers
blanchet
parents: 39342
diff changeset
   864
           | Const (x0 as (@{const_name Ex}, T0))
37704
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   865
             $ (t1 as Abs (s1, T1, t1')) =>
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   866
             (case sn of
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   867
                Plus => do_quantifier x0 s1 T1 t1'
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   868
              | Minus =>
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   869
                (* FIXME: Move elsewhere *)
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   870
                do_term (@{const Not}
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   871
                         $ (HOLogic.eq_const (domain_type T0) $ t1
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   872
                            $ Abs (Name.uu, T1, @{const False}))) accum)
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38795
diff changeset
   873
           | Const (x as (@{const_name HOL.eq}, _)) $ t1 $ t2 =>
37704
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   874
             do_equals x t1 t2
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   875
           | Const (@{const_name Let}, _) $ t1 $ t2 =>
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   876
             do_formula sn (betapply (t2, t1)) accum
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   877
           | (t0 as Const (s0, _)) $ t1 $ t2 =>
38199
8a9cace789d6 deal correctly with Pure.conjunction in mono check
blanchet
parents: 38190
diff changeset
   878
             if s0 = @{const_name "==>"} orelse
8a9cace789d6 deal correctly with Pure.conjunction in mono check
blanchet
parents: 38190
diff changeset
   879
                s0 = @{const_name Pure.conjunction} orelse
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   880
                s0 = @{const_name HOL.conj} orelse
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   881
                s0 = @{const_name HOL.disj} orelse
38786
e46e7a9cb622 formerly unnamed infix impliciation now named HOL.implies
haftmann
parents: 38199
diff changeset
   882
                s0 = @{const_name HOL.implies} then
37704
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   883
               let
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   884
                 val impl = (s0 = @{const_name "==>"} orelse
38786
e46e7a9cb622 formerly unnamed infix impliciation now named HOL.implies
haftmann
parents: 38199
diff changeset
   885
                             s0 = @{const_name HOL.implies})
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   886
                 val (m1, accum) =
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   887
                   do_formula (sn |> impl ? negate_sign) t1 accum
37704
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   888
                 val (m2, accum) = do_formula sn t2 accum
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   889
               in
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   890
                 (MApp (MApp (MRaw (t0, mtype_for (fastype_of t0)), m1), m2),
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   891
                  accum)
39316
b6c4385ab400 change defaults of Auto Nitpick so that it consumes less resources (time and Kodkod threads)
blanchet
parents: 38864
diff changeset
   892
               end
37704
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   893
             else
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   894
               do_term t accum
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   895
           | _ => do_term t accum)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   896
        end
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   897
        |> tap (fn (m, _) =>
38179
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   898
                   trace_msg (fn () => "\<Gamma> \<turnstile> " ^
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   899
                                       string_for_mterm ctxt m ^ " : o\<^sup>" ^
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   900
                                       string_for_sign sn))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   901
  in do_formula end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   902
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   903
(* The harmless axiom optimization below is somewhat too aggressive in the face
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   904
   of (rather peculiar) user-defined axioms. *)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   905
val harmless_consts =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   906
  [@{const_name ord_class.less}, @{const_name ord_class.less_eq}]
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   907
val bounteous_consts = [@{const_name bisim}]
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   908
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   909
fun is_harmless_axiom ({no_harmless = true, ...} : mdata) _ = false
39359
6f49c7fbb1b1 remove "fast_descs" option from Nitpick;
blanchet
parents: 39345
diff changeset
   910
  | is_harmless_axiom {hol_ctxt = {thy, stds, ...}, ...} t =
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   911
    Term.add_consts t []
39359
6f49c7fbb1b1 remove "fast_descs" option from Nitpick;
blanchet
parents: 39345
diff changeset
   912
    |> filter_out (is_built_in_const thy stds)
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   913
    |> (forall (member (op =) harmless_consts o original_name o fst) orf
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   914
        exists (member (op =) bounteous_consts o fst))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   915
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   916
fun consider_nondefinitional_axiom mdata t =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   917
  if is_harmless_axiom mdata t then pair (MRaw (t, dummy_M))
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   918
  else consider_general_formula mdata Plus t
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   919
37256
0dca1ec52999 thread along context instead of theory for typedef lookup
blanchet
parents: 36385
diff changeset
   920
fun consider_definitional_axiom (mdata as {hol_ctxt = {ctxt, ...}, ...}) t =
0dca1ec52999 thread along context instead of theory for typedef lookup
blanchet
parents: 36385
diff changeset
   921
  if not (is_constr_pattern_formula ctxt t) then
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   922
    consider_nondefinitional_axiom mdata t
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   923
  else if is_harmless_axiom mdata t then
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   924
    pair (MRaw (t, dummy_M))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   925
  else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   926
    let
35672
ff484d4f2e14 more work on Nitpick's finite sets
blanchet
parents: 35671
diff changeset
   927
      val mtype_for = fresh_mtype_for_type mdata false
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   928
      val do_term = consider_term mdata
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   929
      fun do_all quant_t abs_s abs_T body_t accum =
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   930
        let
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   931
          val abs_M = mtype_for abs_T
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   932
          val (body_m, accum) =
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   933
            accum |>> push_bound abs_T abs_M |> do_formula body_t
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   934
          val body_M = mtype_of_mterm body_m
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   935
        in
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   936
          (MApp (MRaw (quant_t, MFun (MFun (abs_M, A Gen, body_M), A Gen,
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   937
                       body_M)),
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
   938
                 MAbs (abs_s, abs_T, abs_M, A Gen, body_m)),
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   939
           accum |>> pop_bound)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   940
        end
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   941
      and do_conjunction t0 t1 t2 accum =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   942
        let
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   943
          val (m1, accum) = do_formula t1 accum
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   944
          val (m2, accum) = do_formula t2 accum
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   945
        in
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   946
          (MApp (MApp (MRaw (t0, mtype_for (fastype_of t0)), m1), m2), accum)
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   947
        end
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   948
      and do_implies t0 t1 t2 accum =
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   949
        let
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   950
          val (m1, accum) = do_term t1 accum
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   951
          val (m2, accum) = do_formula t2 accum
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   952
        in
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   953
          (MApp (MApp (MRaw (t0, mtype_for (fastype_of t0)), m1), m2), accum)
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   954
        end
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
   955
      and do_formula t accum =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   956
          case t of
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   957
            (t0 as Const (@{const_name all}, _)) $ Abs (s1, T1, t1) =>
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   958
            do_all t0 s1 T1 t1 accum
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   959
          | @{const Trueprop} $ t1 =>
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   960
            let val (m1, accum) = do_formula t1 accum in
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   961
              (MApp (MRaw (@{const Trueprop}, mtype_for (bool_T --> prop_T)),
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   962
                     m1), accum)
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   963
            end
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   964
          | Const (x as (@{const_name "=="}, _)) $ t1 $ t2 =>
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   965
            consider_general_equals mdata true x t1 t2 accum
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   966
          | (t0 as @{const "==>"}) $ t1 $ t2 => do_implies t0 t1 t2 accum
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   967
          | (t0 as @{const Pure.conjunction}) $ t1 $ t2 =>
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   968
            do_conjunction t0 t1 t2 accum
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   969
          | (t0 as Const (@{const_name All}, _)) $ Abs (s0, T1, t1) =>
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   970
            do_all t0 s0 T1 t1 accum
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38795
diff changeset
   971
          | Const (x as (@{const_name HOL.eq}, _)) $ t1 $ t2 =>
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   972
            consider_general_equals mdata true x t1 t2 accum
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   973
          | (t0 as @{const HOL.conj}) $ t1 $ t2 => do_conjunction t0 t1 t2 accum
38786
e46e7a9cb622 formerly unnamed infix impliciation now named HOL.implies
haftmann
parents: 38199
diff changeset
   974
          | (t0 as @{const HOL.implies}) $ t1 $ t2 => do_implies t0 t1 t2 accum
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
   975
          | _ => raise TERM ("Nitpick_Mono.consider_definitional_axiom.\
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   976
                             \do_formula", [t])
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   977
    in do_formula t end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   978
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   979
fun string_for_mtype_of_term ctxt lits t M =
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   980
  Syntax.string_of_term ctxt t ^ " : " ^ string_for_mtype (resolve_mtype lits M)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   981
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35384
diff changeset
   982
fun print_mtype_context ctxt lits ({frees, consts, ...} : mtype_context) =
38179
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   983
  trace_msg (fn () =>
37476
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
   984
      map (fn (x, M) => string_for_mtype_of_term ctxt lits (Free x) M) frees @
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
   985
      map (fn (x, M) => string_for_mtype_of_term ctxt lits (Const x) M) consts
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
   986
      |> cat_lines)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   987
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   988
fun amass f t (ms, accum) =
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   989
  let val (m, accum) = f t accum in (m :: ms, accum) end
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
   990
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   991
fun infer which no_harmless (hol_ctxt as {ctxt, ...}) binarize alpha_T
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   992
          (nondef_ts, def_ts) =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   993
  let
38179
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   994
    val _ = trace_msg (fn () => "****** " ^ which ^ " analysis: " ^
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   995
                                string_for_mtype MAlpha ^ " is " ^
7207321df8af make tracing monotonicity easier
blanchet
parents: 37704
diff changeset
   996
                                Syntax.string_of_typ ctxt alpha_T)
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   997
    val mdata as {max_fresh, constr_mcache, ...} =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
   998
      initial_mdata hol_ctxt binarize no_harmless alpha_T
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
   999
    val accum = (initial_gamma, ([], [], []))
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
  1000
    val (nondef_ms, accum) =
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
  1001
      ([], accum) |> amass (consider_general_formula mdata Plus) (hd nondef_ts)
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
  1002
                  |> fold (amass (consider_nondefinitional_axiom mdata))
35386
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
  1003
                          (tl nondef_ts)
45a4e19d3ebd more work on the new monotonicity stuff in Nitpick
blanchet
parents: 35385
diff changeset
  1004
    val (def_ms, (gamma, cset)) =
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
  1005
      ([], accum) |> fold (amass (consider_definitional_axiom mdata)) def_ts
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1006
  in
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1007
    case solve (!max_fresh) cset of
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
  1008
      SOME lits => (print_mtype_context ctxt lits gamma;
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
  1009
                    SOME (lits, (nondef_ms, def_ms), !constr_mcache))
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
  1010
    | _ => NONE
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1011
  end
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1012
  handle UNSOLVABLE () => NONE
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1013
       | MTYPE (loc, Ms, Ts) =>
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
  1014
         raise BAD (loc, commas (map string_for_mtype Ms @
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
  1015
                                 map (Syntax.string_of_typ ctxt) Ts))
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
  1016
       | MTERM (loc, ms) =>
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
  1017
         raise BAD (loc, commas (map (string_for_mterm ctxt) ms))
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
  1018
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
  1019
val formulas_monotonic = is_some oooo infer "Monotonicity" false
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
  1020
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
  1021
fun fin_fun_constr T1 T2 =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
  1022
  (@{const_name FinFun}, (T1 --> T2) --> Type (@{type_name fin_fun}, [T1, T2]))
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
  1023
39359
6f49c7fbb1b1 remove "fast_descs" option from Nitpick;
blanchet
parents: 39345
diff changeset
  1024
fun finitize_funs (hol_ctxt as {thy, ctxt, stds, constr_cache, ...})
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
  1025
                  binarize finitizes alpha_T tsp =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
  1026
  case infer "Finiteness" true hol_ctxt binarize alpha_T tsp of
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
  1027
    SOME (lits, msp, constr_mtypes) =>
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
  1028
    if forall (curry (op =) Gen o snd) lits then
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1029
      tsp
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1030
    else
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1031
      let
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1032
        fun should_finitize T a =
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1033
          case triple_lookup (type_match thy) finitizes T of
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1034
            SOME (SOME false) => false
40985
8b870370c26f started generalizing monotonicity code to accommodate new calculus
blanchet
parents: 40722
diff changeset
  1035
          | _ => resolve_annotation_atom lits a = A Fls
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1036
        fun type_from_mtype T M =
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1037
          case (M, T) of
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1038
            (MAlpha, _) => T
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1039
          | (MFun (M1, a, M2), Type (@{type_name fun}, Ts)) =>
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1040
            Type (if should_finitize T a then @{type_name fin_fun}
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1041
                  else @{type_name fun}, map2 type_from_mtype Ts [M1, M2])
38190
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 38179
diff changeset
  1042
          | (MPair (M1, M2), Type (@{type_name prod}, Ts)) =>
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 38179
diff changeset
  1043
            Type (@{type_name prod}, map2 type_from_mtype Ts [M1, M2])
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1044
          | (MType _, _) => T
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1045
          | _ => raise MTYPE ("Nitpick_Mono.finitize_funs.type_from_mtype",
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1046
                              [M], [T])
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1047
        fun finitize_constr (x as (s, T)) =
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1048
          (s, case AList.lookup (op =) constr_mtypes x of
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1049
                SOME M => type_from_mtype T M
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1050
              | NONE => T)
37267
5c47d633c84d fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents: 37260
diff changeset
  1051
        fun term_from_mterm new_Ts old_Ts m =
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1052
          case m of
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1053
            MRaw (t, M) =>
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1054
            let
37267
5c47d633c84d fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents: 37260
diff changeset
  1055
              val T = fastype_of1 (old_Ts, t)
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1056
              val T' = type_from_mtype T M
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1057
            in
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1058
              case t of
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1059
                Const (x as (s, _)) =>
37476
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
  1060
                if s = @{const_name finite} then
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1061
                  case domain_type T' of
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1062
                    set_T' as Type (@{type_name fin_fun}, _) =>
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1063
                    Abs (Name.uu, set_T', @{const True})
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1064
                  | _ => Const (s, T')
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1065
                else if s = @{const_name "=="} orelse
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38795
diff changeset
  1066
                        s = @{const_name HOL.eq} then
37476
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
  1067
                  let
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
  1068
                    val T =
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
  1069
                      case T' of
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
  1070
                        Type (_, [T1, Type (_, [T2, T3])]) =>
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
  1071
                        T1 --> T2 --> T3
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
  1072
                      | _ => raise TYPE ("Nitpick_Mono.finitize_funs.\
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
  1073
                                         \term_from_mterm", [T, T'], [])
0681e46b4022 optimized code generated for datatype cases + more;
blanchet
parents: 37267
diff changeset
  1074
                  in coerce_term hol_ctxt new_Ts T' T (Const (s, T)) end
39359
6f49c7fbb1b1 remove "fast_descs" option from Nitpick;
blanchet
parents: 39345
diff changeset
  1075
                else if is_built_in_const thy stds x then
37267
5c47d633c84d fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents: 37260
diff changeset
  1076
                  coerce_term hol_ctxt new_Ts T' T t
37256
0dca1ec52999 thread along context instead of theory for typedef lookup
blanchet
parents: 36385
diff changeset
  1077
                else if is_constr ctxt stds x then
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1078
                  Const (finitize_constr x)
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1079
                else if is_sel s then
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1080
                  let
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1081
                    val n = sel_no_from_name s
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1082
                    val x' =
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1083
                      x |> binarized_and_boxed_constr_for_sel hol_ctxt binarize
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1084
                        |> finitize_constr
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1085
                    val x'' =
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1086
                      binarized_and_boxed_nth_sel_for_constr hol_ctxt binarize
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1087
                                                             x' n
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1088
                  in Const x'' end
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1089
                else
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1090
                  Const (s, T')
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1091
              | Free (s, T) => Free (s, type_from_mtype T M)
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1092
              | Bound _ => t
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1093
              | _ => raise MTERM ("Nitpick_Mono.finitize_funs.term_from_mterm",
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1094
                                  [m])
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1095
            end
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1096
          | MApp (m1, m2) =>
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1097
            let
37267
5c47d633c84d fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents: 37260
diff changeset
  1098
              val (t1, t2) = pairself (term_from_mterm new_Ts old_Ts) (m1, m2)
5c47d633c84d fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents: 37260
diff changeset
  1099
              val (T1, T2) = pairself (curry fastype_of1 new_Ts) (t1, t2)
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1100
              val (t1', T2') =
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1101
                case T1 of
39316
b6c4385ab400 change defaults of Auto Nitpick so that it consumes less resources (time and Kodkod threads)
blanchet
parents: 38864
diff changeset
  1102
                  Type (s, [T11, T12]) =>
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1103
                  (if s = @{type_name fin_fun} then
37256
0dca1ec52999 thread along context instead of theory for typedef lookup
blanchet
parents: 36385
diff changeset
  1104
                     select_nth_constr_arg ctxt stds (fin_fun_constr T11 T12) t1
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1105
                                           0 (T11 --> T12)
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1106
                   else
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1107
                     t1, T11)
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1108
                | _ => raise TYPE ("Nitpick_Mono.finitize_funs.term_from_mterm",
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1109
                                   [T1], [])
37267
5c47d633c84d fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents: 37260
diff changeset
  1110
            in betapply (t1', coerce_term hol_ctxt new_Ts T2' T2 t2) end
5c47d633c84d fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents: 37260
diff changeset
  1111
          | MAbs (s, old_T, M, a, m') =>
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1112
            let
37267
5c47d633c84d fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents: 37260
diff changeset
  1113
              val new_T = type_from_mtype old_T M
5c47d633c84d fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents: 37260
diff changeset
  1114
              val t' = term_from_mterm (new_T :: new_Ts) (old_T :: old_Ts) m'
5c47d633c84d fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents: 37260
diff changeset
  1115
              val T' = fastype_of1 (new_T :: new_Ts, t')
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1116
            in
37267
5c47d633c84d fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents: 37260
diff changeset
  1117
              Abs (s, new_T, t')
5c47d633c84d fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents: 37260
diff changeset
  1118
              |> should_finitize (new_T --> T') a
5c47d633c84d fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents: 37260
diff changeset
  1119
                 ? construct_value ctxt stds (fin_fun_constr new_T T') o single
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1120
            end
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1121
      in
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1122
        Unsynchronized.change constr_cache (map (apsnd (map finitize_constr)));
37267
5c47d633c84d fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents: 37260
diff changeset
  1123
        pairself (map (term_from_mterm [] [])) msp
35812
394fe2b064cd solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents: 35672
diff changeset
  1124
      end
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35386
diff changeset
  1125
  | NONE => tsp
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1126
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
  1127
end;