src/HOL/Nitpick_Examples/minipick.ML
author wenzelm
Tue, 29 Sep 2020 11:35:21 +0200
changeset 72328 7cb0c5fbe2d9
parent 72205 bc71db05abe3
child 72329 6255e532aa36
permissions -rw-r--r--
obsolete --- KODKODI is always present via component;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/Nitpick_Examples/minipick.ML
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
     3
    Copyright   2009-2010
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
     4
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
     5
Finite model generation for HOL formulas using Kodkod, minimalistic version.
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
     6
*)
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
     7
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
     8
signature MINIPICK =
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
     9
sig
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    10
  val minipick : Proof.context -> int -> term -> string
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    11
  val minipick_expect : Proof.context -> string -> int -> term -> unit
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    12
end;
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    13
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    14
structure Minipick : MINIPICK =
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    15
struct
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    16
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    17
open Kodkod
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    18
open Nitpick_Util
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    19
open Nitpick_HOL
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    20
open Nitpick_Peephole
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    21
open Nitpick_Kodkod
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    22
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    23
datatype rep =
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    24
  S_Rep |
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    25
  R_Rep of bool
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    26
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
    27
fun check_type ctxt raw_infinite (Type (\<^type_name>\<open>fun\<close>, Ts)) =
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    28
    List.app (check_type ctxt raw_infinite) Ts
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
    29
  | check_type ctxt raw_infinite (Type (\<^type_name>\<open>prod\<close>, Ts)) =
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    30
    List.app (check_type ctxt raw_infinite) Ts
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
    31
  | check_type _ _ \<^typ>\<open>bool\<close> = ()
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
    32
  | check_type _ _ (TFree (_, \<^sort>\<open>{}\<close>)) = ()
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
    33
  | check_type _ _ (TFree (_, \<^sort>\<open>HOL.type\<close>)) = ()
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    34
  | check_type ctxt raw_infinite T =
46092
287a3cefc21b ported Minipick to "set"
blanchet
parents: 45128
diff changeset
    35
    if raw_infinite T then
287a3cefc21b ported Minipick to "set"
blanchet
parents: 45128
diff changeset
    36
      ()
287a3cefc21b ported Minipick to "set"
blanchet
parents: 45128
diff changeset
    37
    else
287a3cefc21b ported Minipick to "set"
blanchet
parents: 45128
diff changeset
    38
      error ("Not supported: Type " ^ quote (Syntax.string_of_typ ctxt T) ^ ".")
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    39
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
    40
fun atom_schema_of S_Rep card (Type (\<^type_name>\<open>fun\<close>, [T1, T2])) =
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    41
    replicate_list (card T1) (atom_schema_of S_Rep card T2)
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    42
  | atom_schema_of (R_Rep true) card
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
    43
                   (Type (\<^type_name>\<open>fun\<close>, [T1, \<^typ>\<open>bool\<close>])) =
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    44
    atom_schema_of S_Rep card T1
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
    45
  | atom_schema_of (rep as R_Rep _) card (Type (\<^type_name>\<open>fun\<close>, [T1, T2])) =
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    46
    atom_schema_of S_Rep card T1 @ atom_schema_of rep card T2
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
    47
  | atom_schema_of _ card (Type (\<^type_name>\<open>prod\<close>, Ts)) =
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    48
    maps (atom_schema_of S_Rep card) Ts
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    49
  | atom_schema_of _ card T = [card T]
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    50
val arity_of = length ooo atom_schema_of
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    51
val atom_seqs_of = map (AtomSeq o rpair 0) ooo atom_schema_of
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    52
val atom_seq_product_of = foldl1 Product ooo atom_seqs_of
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    53
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    54
fun index_for_bound_var _ [_] 0 = 0
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    55
  | index_for_bound_var card (_ :: Ts) 0 =
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    56
    index_for_bound_var card Ts 0 + arity_of S_Rep card (hd Ts)
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    57
  | index_for_bound_var card Ts n = index_for_bound_var card (tl Ts) (n - 1)
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    58
fun vars_for_bound_var card R Ts j =
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    59
  map (curry Var 1) (index_seq (index_for_bound_var card Ts j)
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    60
                               (arity_of R card (nth Ts j)))
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    61
val rel_expr_for_bound_var = foldl1 Product oooo vars_for_bound_var
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    62
fun decls_for R card Ts T =
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    63
  map2 (curry DeclOne o pair 1)
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    64
       (index_seq (index_for_bound_var card (T :: Ts) 0)
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    65
                  (arity_of R card (nth (T :: Ts) 0)))
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    66
       (atom_seqs_of R card T)
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    67
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    68
val atom_product = foldl1 Product o map Atom
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    69
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    70
val false_atom_num = 0
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    71
val true_atom_num = 1
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    72
val false_atom = Atom false_atom_num
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    73
val true_atom = Atom true_atom_num
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    74
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    75
fun kodkod_formula_from_term ctxt total card complete concrete frees =
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
    76
  let
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    77
    fun F_from_S_rep (SOME false) r = Not (RelEq (r, false_atom))
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    78
      | F_from_S_rep _ r = RelEq (r, true_atom)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    79
    fun S_rep_from_F NONE f = RelIf (f, true_atom, false_atom)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    80
      | S_rep_from_F (SOME true) f = RelIf (f, true_atom, None)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    81
      | S_rep_from_F (SOME false) f = RelIf (Not f, false_atom, None)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
    82
    fun R_rep_from_S_rep (Type (\<^type_name>\<open>fun\<close>, [T1, T2])) r =
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    83
        if total andalso T2 = bool_T then
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    84
          let
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    85
            val jss = atom_schema_of S_Rep card T1 |> map (rpair 0)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    86
                      |> all_combinations
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    87
          in
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    88
            map2 (fn i => fn js =>
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    89
(*
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    90
                     RelIf (F_from_S_rep NONE (Project (r, [Num i])),
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    91
                            atom_product js, empty_n_ary_rel (length js))
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    92
*)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    93
                     Join (Project (r, [Num i]),
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    94
                           atom_product (false_atom_num :: js))
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    95
                 ) (index_seq 0 (length jss)) jss
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    96
            |> foldl1 Union
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    97
          end
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    98
        else
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
    99
          let
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   100
            val jss = atom_schema_of S_Rep card T1 |> map (rpair 0)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   101
                      |> all_combinations
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   102
            val arity2 = arity_of S_Rep card T2
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   103
          in
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   104
            map2 (fn i => fn js =>
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   105
                     Product (atom_product js,
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   106
                              Project (r, num_seq (i * arity2) arity2)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   107
                              |> R_rep_from_S_rep T2))
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   108
                 (index_seq 0 (length jss)) jss
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   109
            |> foldl1 Union
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   110
          end
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   111
      | R_rep_from_S_rep _ r = r
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   112
    fun S_rep_from_R_rep Ts (T as Type (\<^type_name>\<open>fun\<close>, _)) r =
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   113
        Comprehension (decls_for S_Rep card Ts T,
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   114
            RelEq (R_rep_from_S_rep T
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   115
                       (rel_expr_for_bound_var card S_Rep (T :: Ts) 0), r))
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   116
      | S_rep_from_R_rep _ _ r = r
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   117
    fun partial_eq pos Ts (Type (\<^type_name>\<open>fun\<close>, [T1, T2])) t1 t2 =
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   118
        HOLogic.mk_all ("x", T1,
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   119
                        HOLogic.eq_const T2 $ (incr_boundvars 1 t1 $ Bound 0)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   120
                        $ (incr_boundvars 1 t2 $ Bound 0))
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   121
        |> to_F (SOME pos) Ts
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   122
      | partial_eq pos Ts T t1 t2 =
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   123
        if pos andalso not (concrete T) then
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   124
          False
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   125
        else
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 55208
diff changeset
   126
          (t1, t2) |> apply2 (to_R_rep Ts)
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   127
                   |> (if pos then Some o Intersect else Lone o Union)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   128
    and to_F pos Ts t =
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   129
      (case t of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   130
         \<^const>\<open>Not\<close> $ t1 => Not (to_F (Option.map not pos) Ts t1)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   131
       | \<^const>\<open>False\<close> => False
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   132
       | \<^const>\<open>True\<close> => True
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   133
       | Const (\<^const_name>\<open>All\<close>, _) $ Abs (_, T, t') =>
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   134
         if pos = SOME true andalso not (complete T) then False
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   135
         else All (decls_for S_Rep card Ts T, to_F pos (T :: Ts) t')
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   136
       | (t0 as Const (\<^const_name>\<open>All\<close>, _)) $ t1 =>
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   137
         to_F pos Ts (t0 $ eta_expand Ts t1 1)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   138
       | Const (\<^const_name>\<open>Ex\<close>, _) $ Abs (_, T, t') =>
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   139
         if pos = SOME false andalso not (complete T) then True
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   140
         else Exist (decls_for S_Rep card Ts T, to_F pos (T :: Ts) t')
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   141
       | (t0 as Const (\<^const_name>\<open>Ex\<close>, _)) $ t1 =>
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   142
         to_F pos Ts (t0 $ eta_expand Ts t1 1)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   143
       | Const (\<^const_name>\<open>HOL.eq\<close>, Type (_, [T, _])) $ t1 $ t2 =>
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   144
         (case pos of
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   145
            NONE => RelEq (to_R_rep Ts t1, to_R_rep Ts t2)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   146
          | SOME pos => partial_eq pos Ts T t1 t2)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   147
       | Const (\<^const_name>\<open>ord_class.less_eq\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   148
                Type (\<^type_name>\<open>fun\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   149
                      [Type (\<^type_name>\<open>fun\<close>, [T', \<^typ>\<open>bool\<close>]), _]))
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   150
         $ t1 $ t2 =>
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   151
         (case pos of
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   152
            NONE => Subset (to_R_rep Ts t1, to_R_rep Ts t2)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   153
          | SOME true =>
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   154
            Subset (Difference (atom_seq_product_of S_Rep card T',
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   155
                                Join (to_R_rep Ts t1, false_atom)),
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   156
                    Join (to_R_rep Ts t2, true_atom))
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   157
          | SOME false =>
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   158
            Subset (Join (to_R_rep Ts t1, true_atom),
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   159
                    Difference (atom_seq_product_of S_Rep card T',
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   160
                                Join (to_R_rep Ts t2, false_atom))))
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   161
       | \<^const>\<open>HOL.conj\<close> $ t1 $ t2 => And (to_F pos Ts t1, to_F pos Ts t2)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   162
       | \<^const>\<open>HOL.disj\<close> $ t1 $ t2 => Or (to_F pos Ts t1, to_F pos Ts t2)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   163
       | \<^const>\<open>HOL.implies\<close> $ t1 $ t2 =>
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   164
         Implies (to_F (Option.map not pos) Ts t1, to_F pos Ts t2)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   165
       | Const (\<^const_name>\<open>Set.member\<close>, _) $ t1 $ t2 => to_F pos Ts (t2 $ t1)
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   166
       | t1 $ t2 =>
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   167
         (case pos of
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   168
            NONE => Subset (to_S_rep Ts t2, to_R_rep Ts t1)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   169
          | SOME pos =>
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   170
            let
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   171
              val kt1 = to_R_rep Ts t1
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   172
              val kt2 = to_S_rep Ts t2
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   173
              val kT = atom_seq_product_of S_Rep card (fastype_of1 (Ts, t2))
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   174
            in
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   175
              if pos then
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   176
                Not (Subset (kt2, Difference (kT, Join (kt1, true_atom))))
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   177
              else
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   178
                Subset (kt2, Difference (kT, Join (kt1, false_atom)))
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   179
            end)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   180
       | _ => raise SAME ())
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   181
      handle SAME () => F_from_S_rep pos (to_R_rep Ts t)
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   182
    and to_S_rep Ts t =
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   183
      case t of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   184
        Const (\<^const_name>\<open>Pair\<close>, _) $ t1 $ t2 =>
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   185
        Product (to_S_rep Ts t1, to_S_rep Ts t2)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   186
      | Const (\<^const_name>\<open>Pair\<close>, _) $ _ => to_S_rep Ts (eta_expand Ts t 1)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   187
      | Const (\<^const_name>\<open>Pair\<close>, _) => to_S_rep Ts (eta_expand Ts t 2)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   188
      | Const (\<^const_name>\<open>fst\<close>, _) $ t1 =>
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   189
        let val fst_arity = arity_of S_Rep card (fastype_of1 (Ts, t)) in
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   190
          Project (to_S_rep Ts t1, num_seq 0 fst_arity)
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   191
        end
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   192
      | Const (\<^const_name>\<open>fst\<close>, _) => to_S_rep Ts (eta_expand Ts t 1)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   193
      | Const (\<^const_name>\<open>snd\<close>, _) $ t1 =>
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   194
        let
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   195
          val pair_arity = arity_of S_Rep card (fastype_of1 (Ts, t1))
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   196
          val snd_arity = arity_of S_Rep card (fastype_of1 (Ts, t))
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   197
          val fst_arity = pair_arity - snd_arity
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   198
        in Project (to_S_rep Ts t1, num_seq fst_arity snd_arity) end
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   199
      | Const (\<^const_name>\<open>snd\<close>, _) => to_S_rep Ts (eta_expand Ts t 1)
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   200
      | Bound j => rel_expr_for_bound_var card S_Rep Ts j
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   201
      | _ => S_rep_from_R_rep Ts (fastype_of1 (Ts, t)) (to_R_rep Ts t)
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   202
    and partial_set_op swap1 swap2 op1 op2 Ts t1 t2 =
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   203
      let
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   204
        val kt1 = to_R_rep Ts t1
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   205
        val kt2 = to_R_rep Ts t2
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   206
        val (a11, a21) = (false_atom, true_atom) |> swap1 ? swap
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   207
        val (a12, a22) = (false_atom, true_atom) |> swap2 ? swap
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   208
      in
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   209
        Union (Product (op1 (Join (kt1, a11), Join (kt2, a12)), true_atom),
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   210
               Product (op2 (Join (kt1, a21), Join (kt2, a22)), false_atom))
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   211
      end
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   212
    and to_R_rep Ts t =
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   213
      (case t of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   214
         \<^const>\<open>Not\<close> => to_R_rep Ts (eta_expand Ts t 1)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   215
       | Const (\<^const_name>\<open>All\<close>, _) => to_R_rep Ts (eta_expand Ts t 1)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   216
       | Const (\<^const_name>\<open>Ex\<close>, _) => to_R_rep Ts (eta_expand Ts t 1)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   217
       | Const (\<^const_name>\<open>HOL.eq\<close>, _) $ _ => to_R_rep Ts (eta_expand Ts t 1)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   218
       | Const (\<^const_name>\<open>HOL.eq\<close>, _) => to_R_rep Ts (eta_expand Ts t 2)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   219
       | Const (\<^const_name>\<open>ord_class.less_eq\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   220
                Type (\<^type_name>\<open>fun\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   221
                      [Type (\<^type_name>\<open>fun\<close>, [_, \<^typ>\<open>bool\<close>]), _])) $ _ =>
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   222
         to_R_rep Ts (eta_expand Ts t 1)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   223
       | Const (\<^const_name>\<open>ord_class.less_eq\<close>, _) =>
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   224
         to_R_rep Ts (eta_expand Ts t 2)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   225
       | \<^const>\<open>HOL.conj\<close> $ _ => to_R_rep Ts (eta_expand Ts t 1)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   226
       | \<^const>\<open>HOL.conj\<close> => to_R_rep Ts (eta_expand Ts t 2)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   227
       | \<^const>\<open>HOL.disj\<close> $ _ => to_R_rep Ts (eta_expand Ts t 1)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   228
       | \<^const>\<open>HOL.disj\<close> => to_R_rep Ts (eta_expand Ts t 2)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   229
       | \<^const>\<open>HOL.implies\<close> $ _ => to_R_rep Ts (eta_expand Ts t 1)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   230
       | \<^const>\<open>HOL.implies\<close> => to_R_rep Ts (eta_expand Ts t 2)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   231
       | Const (\<^const_name>\<open>Set.member\<close>, _) $ _ =>
46092
287a3cefc21b ported Minipick to "set"
blanchet
parents: 45128
diff changeset
   232
         to_R_rep Ts (eta_expand Ts t 1)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   233
       | Const (\<^const_name>\<open>Set.member\<close>, _) => to_R_rep Ts (eta_expand Ts t 2)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   234
       | Const (\<^const_name>\<open>Collect\<close>, _) $ t' => to_R_rep Ts t'
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   235
       | Const (\<^const_name>\<open>Collect\<close>, _) => to_R_rep Ts (eta_expand Ts t 1)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   236
       | Const (\<^const_name>\<open>bot_class.bot\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   237
                T as Type (\<^type_name>\<open>fun\<close>, [T', \<^typ>\<open>bool\<close>])) =>
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   238
         if total then empty_n_ary_rel (arity_of (R_Rep total) card T)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   239
         else Product (atom_seq_product_of (R_Rep total) card T', false_atom)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   240
       | Const (\<^const_name>\<open>top_class.top\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   241
                T as Type (\<^type_name>\<open>fun\<close>, [T', \<^typ>\<open>bool\<close>])) =>
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   242
         if total then atom_seq_product_of (R_Rep total) card T
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   243
         else Product (atom_seq_product_of (R_Rep total) card T', true_atom)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   244
       | Const (\<^const_name>\<open>insert\<close>, Type (_, [T, _])) $ t1 $ t2 =>
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   245
         if total then
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   246
           Union (to_S_rep Ts t1, to_R_rep Ts t2)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   247
         else
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   248
           let
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   249
             val kt1 = to_S_rep Ts t1
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   250
             val kt2 = to_R_rep Ts t2
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   251
           in
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   252
             RelIf (Some kt1,
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   253
                    if arity_of S_Rep card T = 1 then
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   254
                      Override (kt2, Product (kt1, true_atom))
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   255
                    else
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   256
                      Union (Difference (kt2, Product (kt1, false_atom)),
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   257
                             Product (kt1, true_atom)),
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   258
                    Difference (kt2, Product (atom_seq_product_of S_Rep card T,
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   259
                                              false_atom)))
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   260
           end
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   261
       | Const (\<^const_name>\<open>insert\<close>, _) $ _ => to_R_rep Ts (eta_expand Ts t 1)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   262
       | Const (\<^const_name>\<open>insert\<close>, _) => to_R_rep Ts (eta_expand Ts t 2)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   263
       | Const (\<^const_name>\<open>trancl\<close>,
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   264
                Type (_, [Type (_, [Type (_, [T', _]), _]), _])) $ t1 =>
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   265
         if arity_of S_Rep card T' = 1 then
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   266
           if total then
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   267
             Closure (to_R_rep Ts t1)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   268
           else
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   269
             let
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   270
               val kt1 = to_R_rep Ts t1
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   271
               val true_core_kt = Closure (Join (kt1, true_atom))
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   272
               val kTx =
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   273
                 atom_seq_product_of S_Rep card (HOLogic.mk_prodT (`I T'))
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   274
               val false_mantle_kt =
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   275
                 Difference (kTx,
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   276
                     Closure (Difference (kTx, Join (kt1, false_atom))))
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   277
             in
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   278
               Union (Product (Difference (false_mantle_kt, true_core_kt),
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   279
                               false_atom),
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   280
                      Product (true_core_kt, true_atom))
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   281
             end
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   282
         else
46092
287a3cefc21b ported Minipick to "set"
blanchet
parents: 45128
diff changeset
   283
           error "Not supported: Transitive closure for function or pair type."
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   284
       | Const (\<^const_name>\<open>trancl\<close>, _) => to_R_rep Ts (eta_expand Ts t 1)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   285
       | Const (\<^const_name>\<open>inf_class.inf\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   286
                Type (\<^type_name>\<open>fun\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   287
                      [Type (\<^type_name>\<open>fun\<close>, [_, \<^typ>\<open>bool\<close>]), _]))
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   288
         $ t1 $ t2 =>
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   289
         if total then Intersect (to_R_rep Ts t1, to_R_rep Ts t2)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   290
         else partial_set_op true true Intersect Union Ts t1 t2
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   291
       | Const (\<^const_name>\<open>inf_class.inf\<close>, _) $ _ =>
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   292
         to_R_rep Ts (eta_expand Ts t 1)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   293
       | Const (\<^const_name>\<open>inf_class.inf\<close>, _) =>
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   294
         to_R_rep Ts (eta_expand Ts t 2)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   295
       | Const (\<^const_name>\<open>sup_class.sup\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   296
                Type (\<^type_name>\<open>fun\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   297
                      [Type (\<^type_name>\<open>fun\<close>, [_, \<^typ>\<open>bool\<close>]), _]))
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   298
         $ t1 $ t2 =>
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   299
         if total then Union (to_R_rep Ts t1, to_R_rep Ts t2)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   300
         else partial_set_op true true Union Intersect Ts t1 t2
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   301
       | Const (\<^const_name>\<open>sup_class.sup\<close>, _) $ _ =>
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   302
         to_R_rep Ts (eta_expand Ts t 1)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   303
       | Const (\<^const_name>\<open>sup_class.sup\<close>, _) =>
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   304
         to_R_rep Ts (eta_expand Ts t 2)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   305
       | Const (\<^const_name>\<open>minus_class.minus\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   306
                Type (\<^type_name>\<open>fun\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   307
                      [Type (\<^type_name>\<open>fun\<close>, [_, \<^typ>\<open>bool\<close>]), _]))
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   308
         $ t1 $ t2 =>
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   309
         if total then Difference (to_R_rep Ts t1, to_R_rep Ts t2)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   310
         else partial_set_op true false Intersect Union Ts t1 t2
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   311
       | Const (\<^const_name>\<open>minus_class.minus\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   312
                Type (\<^type_name>\<open>fun\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   313
                      [Type (\<^type_name>\<open>fun\<close>, [_, \<^typ>\<open>bool\<close>]), _])) $ _ =>
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   314
         to_R_rep Ts (eta_expand Ts t 1)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   315
       | Const (\<^const_name>\<open>minus_class.minus\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   316
                Type (\<^type_name>\<open>fun\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   317
                      [Type (\<^type_name>\<open>fun\<close>, [_, \<^typ>\<open>bool\<close>]), _])) =>
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   318
         to_R_rep Ts (eta_expand Ts t 2)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   319
       | Const (\<^const_name>\<open>Pair\<close>, _) $ _ $ _ => to_S_rep Ts t
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   320
       | Const (\<^const_name>\<open>Pair\<close>, _) $ _ => to_S_rep Ts t
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   321
       | Const (\<^const_name>\<open>Pair\<close>, _) => to_S_rep Ts t
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   322
       | Const (\<^const_name>\<open>fst\<close>, _) $ _ => raise SAME ()
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   323
       | Const (\<^const_name>\<open>fst\<close>, _) => raise SAME ()
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   324
       | Const (\<^const_name>\<open>snd\<close>, _) $ _ => raise SAME ()
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   325
       | Const (\<^const_name>\<open>snd\<close>, _) => raise SAME ()
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   326
       | \<^const>\<open>False\<close> => false_atom
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   327
       | \<^const>\<open>True\<close> => true_atom
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   328
       | Free (x as (_, T)) =>
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   329
         Rel (arity_of (R_Rep total) card T, find_index (curry (op =) x) frees)
46092
287a3cefc21b ported Minipick to "set"
blanchet
parents: 45128
diff changeset
   330
       | Term.Var _ => error "Not supported: Schematic variables."
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   331
       | Bound _ => raise SAME ()
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   332
       | Abs (_, T, t') =>
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   333
         (case (total, fastype_of1 (T :: Ts, t')) of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   334
            (true, \<^typ>\<open>bool\<close>) =>
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   335
            Comprehension (decls_for S_Rep card Ts T, to_F NONE (T :: Ts) t')
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   336
          | (_, T') =>
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   337
            Comprehension (decls_for S_Rep card Ts T @
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   338
                           decls_for (R_Rep total) card (T :: Ts) T',
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   339
                           Subset (rel_expr_for_bound_var card (R_Rep total)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   340
                                                          (T' :: T :: Ts) 0,
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   341
                                   to_R_rep (T :: Ts) t')))
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   342
       | t1 $ t2 =>
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   343
         (case fastype_of1 (Ts, t) of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   344
            \<^typ>\<open>bool\<close> =>
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   345
            if total then
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   346
              S_rep_from_F NONE (to_F NONE Ts t)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   347
            else
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   348
              RelIf (to_F (SOME true) Ts t, true_atom,
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   349
                     RelIf (Not (to_F (SOME false) Ts t), false_atom,
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   350
                     None))
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   351
          | T =>
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   352
            let val T2 = fastype_of1 (Ts, t2) in
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   353
              case arity_of S_Rep card T2 of
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   354
                1 => Join (to_S_rep Ts t2, to_R_rep Ts t1)
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   355
              | arity2 =>
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   356
                let val res_arity = arity_of (R_Rep total) card T in
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   357
                  Project (Intersect
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   358
                      (Product (to_S_rep Ts t2,
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   359
                                atom_seq_product_of (R_Rep total) card T),
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   360
                       to_R_rep Ts t1),
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   361
                      num_seq arity2 res_arity)
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   362
                end
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   363
            end)
46092
287a3cefc21b ported Minipick to "set"
blanchet
parents: 45128
diff changeset
   364
       | _ => error ("Not supported: Term " ^
287a3cefc21b ported Minipick to "set"
blanchet
parents: 45128
diff changeset
   365
                     quote (Syntax.string_of_term ctxt t) ^ "."))
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   366
      handle SAME () => R_rep_from_S_rep (fastype_of1 (Ts, t)) (to_S_rep Ts t)
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   367
  in to_F (if total then NONE else SOME true) [] end
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   368
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   369
fun bound_for_free total card i (s, T) =
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   370
  let val js = atom_schema_of (R_Rep total) card T in
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   371
    ([((length js, i), s)],
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   372
     [TupleSet [], atom_schema_of (R_Rep total) card T |> map (rpair 0)
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   373
                   |> tuple_set_from_atom_schema])
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   374
  end
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   375
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   376
fun declarative_axiom_for_rel_expr total card Ts
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   377
                                   (Type (\<^type_name>\<open>fun\<close>, [T1, T2])) r =
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   378
    if total andalso body_type T2 = bool_T then
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   379
      True
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   380
    else
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   381
      All (decls_for S_Rep card Ts T1,
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   382
           declarative_axiom_for_rel_expr total card (T1 :: Ts) T2
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   383
               (List.foldl Join r (vars_for_bound_var card S_Rep (T1 :: Ts) 0)))
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   384
  | declarative_axiom_for_rel_expr total _ _ _ r =
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   385
    (if total then One else Lone) r
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   386
fun declarative_axiom_for_free total card i (_, T) =
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   387
  declarative_axiom_for_rel_expr total card [] T
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   388
      (Rel (arity_of (R_Rep total) card T, i))
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   389
46092
287a3cefc21b ported Minipick to "set"
blanchet
parents: 45128
diff changeset
   390
(* Hack to make the old code work as is with sets. *)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   391
fun unsetify_type (Type (\<^type_name>\<open>set\<close>, [T])) = unsetify_type T --> bool_T
46092
287a3cefc21b ported Minipick to "set"
blanchet
parents: 45128
diff changeset
   392
  | unsetify_type (Type (s, Ts)) = Type (s, map unsetify_type Ts)
287a3cefc21b ported Minipick to "set"
blanchet
parents: 45128
diff changeset
   393
  | unsetify_type T = T
287a3cefc21b ported Minipick to "set"
blanchet
parents: 45128
diff changeset
   394
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   395
fun kodkod_problem_from_term ctxt total raw_card raw_infinite t =
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   396
  let
45128
5af3a3203a76 discontinued obsolete alias structure ProofContext;
wenzelm
parents: 45062
diff changeset
   397
    val thy = Proof_Context.theory_of ctxt
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   398
    fun card (Type (\<^type_name>\<open>fun\<close>, [T1, T2])) =
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   399
        reasonable_power (card T2) (card T1)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   400
      | card (Type (\<^type_name>\<open>prod\<close>, [T1, T2])) = card T1 * card T2
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   401
      | card \<^typ>\<open>bool\<close> = 2
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   402
      | card T = Int.max (1, raw_card T)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   403
    fun complete (Type (\<^type_name>\<open>fun\<close>, [T1, T2])) =
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   404
        concrete T1 andalso complete T2
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   405
      | complete (Type (\<^type_name>\<open>prod\<close>, Ts)) = forall complete Ts
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   406
      | complete T = not (raw_infinite T)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   407
    and concrete (Type (\<^type_name>\<open>fun\<close>, [T1, T2])) =
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   408
        complete T1 andalso concrete T2
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   409
      | concrete (Type (\<^type_name>\<open>prod\<close>, Ts)) = forall concrete Ts
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   410
      | concrete _ = true
46092
287a3cefc21b ported Minipick to "set"
blanchet
parents: 45128
diff changeset
   411
    val neg_t =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   412
      \<^const>\<open>Not\<close> $ Object_Logic.atomize_term ctxt t
46092
287a3cefc21b ported Minipick to "set"
blanchet
parents: 45128
diff changeset
   413
      |> map_types unsetify_type
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   414
    val _ = fold_types (K o check_type ctxt raw_infinite) neg_t ()
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   415
    val frees = Term.add_frees neg_t []
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   416
    val bounds =
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   417
      map2 (bound_for_free total card) (index_seq 0 (length frees)) frees
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   418
    val declarative_axioms =
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   419
      map2 (declarative_axiom_for_free total card)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   420
           (index_seq 0 (length frees)) frees
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   421
    val formula =
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   422
      neg_t |> kodkod_formula_from_term ctxt total card complete concrete frees 
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   423
            |> fold_rev (curry And) declarative_axioms
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   424
    val univ_card = univ_card 0 0 0 bounds formula
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   425
  in
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   426
    {comment = "", settings = [], univ_card = univ_card, tuple_assigns = [],
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   427
     bounds = bounds, int_bounds = [], expr_assigns = [], formula = formula}
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   428
  end
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   429
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   430
fun solve_any_kodkod_problem thy problems =
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   431
  let
55199
ba93ef2c0d27 tuned ML file name
blanchet
parents: 54845
diff changeset
   432
    val {debug, overlord, timeout, ...} = Nitpick_Commands.default_params thy []
72190
8009c4b5db5e proper treatment of absolute deadline vs. relative timeout;
wenzelm
parents: 69597
diff changeset
   433
    val deadline = Time.now () + timeout
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   434
    val max_threads = 1
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   435
    val max_solutions = 1
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   436
  in
72190
8009c4b5db5e proper treatment of absolute deadline vs. relative timeout;
wenzelm
parents: 69597
diff changeset
   437
    case solve_any_problem debug overlord deadline max_threads max_solutions
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   438
                           problems of
49026
72dcf53c1ee4 adjust example
blanchet
parents: 46092
diff changeset
   439
      JavaNotFound => "unknown"
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   440
    | Normal ([], _, _) => "none"
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   441
    | Normal _ => "genuine"
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   442
    | TimedOut _ => "unknown"
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   443
    | Error (s, _) => error ("Kodkod error: " ^ s)
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   444
  end
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   445
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 59970
diff changeset
   446
val default_raw_infinite = member (op =) [\<^typ>\<open>nat\<close>, \<^typ>\<open>int\<close>]
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   447
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   448
fun minipick ctxt n t =
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   449
  let
45128
5af3a3203a76 discontinued obsolete alias structure ProofContext;
wenzelm
parents: 45062
diff changeset
   450
    val thy = Proof_Context.theory_of ctxt
55199
ba93ef2c0d27 tuned ML file name
blanchet
parents: 54845
diff changeset
   451
    val {total_consts, ...} = Nitpick_Commands.default_params thy []
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   452
    val totals =
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   453
      total_consts |> Option.map single |> the_default [true, false]
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   454
    fun problem_for (total, k) =
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   455
      kodkod_problem_from_term ctxt total (K k) default_raw_infinite t
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   456
  in
55208
11dd3d1da83b compile
blanchet
parents: 55199
diff changeset
   457
    (totals, 1 upto n)
45062
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   458
    |-> map_product pair
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   459
    |> map problem_for
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   460
    |> solve_any_kodkod_problem (Proof_Context.theory_of ctxt)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   461
  end
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   462
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   463
fun minipick_expect ctxt expect n t =
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   464
  if getenv "KODKODI" <> "" then
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   465
    if minipick ctxt n t = expect then ()
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   466
    else error ("\"minipick_expect\" expected " ^ quote expect)
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   467
  else
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   468
    ()
9598cada31b3 first step towards extending Minipick with more translations
blanchet
parents: 45035
diff changeset
   469
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents:
diff changeset
   470
end;