src/HOL/Tools/refute.ML
author urbanc
Fri, 20 Apr 2007 00:28:07 +0200
changeset 22732 5bd1a2a94e1b
parent 22580 d91b4dd651d6
child 22846 fb79144af9a3
permissions -rw-r--r--
declared lemmas true_eqvt and false_eqvt to be equivariant (suggested by samth at ccs.neu.edu)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
     1
(*  Title:      HOL/Tools/refute.ML
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
     2
    ID:         $Id$
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
     3
    Author:     Tjark Weber
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
     4
    Copyright   2003-2007
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
     5
14965
7155b319eafa new SAT solver interface
webertj
parents: 14818
diff changeset
     6
Finite model generation for HOL formulas, using a SAT solver.
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
     7
*)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
     8
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
     9
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    10
(* Declares the 'REFUTE' signature as well as a structure 'Refute'.          *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    11
(* Documentation is available in the Isabelle/Isar theory 'HOL/Refute.thy'.  *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    12
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    13
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    14
signature REFUTE =
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    15
sig
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    16
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    17
  exception REFUTE of string * string
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    18
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    19
(* ------------------------------------------------------------------------- *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
    20
(* Model/interpretation related code (translation HOL -> propositional logic *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    21
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    22
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    23
  type params
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    24
  type interpretation
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    25
  type model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    26
  type arguments
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    27
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    28
  exception MAXVARS_EXCEEDED
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    29
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    30
  val add_interpreter : string -> (theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    31
    (interpretation * model * arguments) option) -> theory -> theory
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    32
  val add_printer     : string -> (theory -> model -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    33
    interpretation -> (int -> bool) -> Term.term option) -> theory -> theory
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    34
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    35
  val interpret : theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    36
    (interpretation * model * arguments)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
    37
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    38
  val print       : theory -> model -> Term.term -> interpretation ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    39
    (int -> bool) -> Term.term
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    40
  val print_model : theory -> model -> (int -> bool) -> string
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    41
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    42
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    43
(* Interface                                                                 *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    44
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    45
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    46
  val set_default_param  : (string * string) -> theory -> theory
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    47
  val get_default_param  : theory -> string -> string option
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    48
  val get_default_params : theory -> (string * string) list
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    49
  val actual_params      : theory -> (string * string) list -> params
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    50
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    51
  val find_model : theory -> params -> Term.term -> bool -> unit
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    52
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    53
  (* tries to find a model for a formula: *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    54
  val satisfy_term   : theory -> (string * string) list -> Term.term -> unit
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    55
  (* tries to find a model that refutes a formula: *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    56
  val refute_term    : theory -> (string * string) list -> Term.term -> unit
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    57
  val refute_subgoal :
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    58
    theory -> (string * string) list -> Thm.thm -> int -> unit
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    59
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    60
  val setup : theory -> theory
22092
ab3dfcef6489 reformatted to 80 chars/line
webertj
parents: 22055
diff changeset
    61
ab3dfcef6489 reformatted to 80 chars/line
webertj
parents: 22055
diff changeset
    62
end;  (* signature REFUTE *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    63
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    64
structure Refute : REFUTE =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    65
struct
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    66
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    67
  open PropLogic;
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    68
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    69
  (* We use 'REFUTE' only for internal error conditions that should    *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    70
  (* never occur in the first place (i.e. errors caused by bugs in our *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    71
  (* code).  Otherwise (e.g. to indicate invalid input data) we use    *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    72
  (* 'error'.                                                          *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    73
  exception REFUTE of string * string;  (* ("in function", "cause") *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    74
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    75
  (* should be raised by an interpreter when more variables would be *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    76
  (* required than allowed by 'maxvars'                              *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    77
  exception MAXVARS_EXCEEDED;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    78
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    79
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    80
(* TREES                                                                     *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    81
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    82
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    83
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    84
(* tree: implements an arbitrarily (but finitely) branching tree as a list   *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    85
(*       of (lists of ...) elements                                          *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    86
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    87
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    88
  datatype 'a tree =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    89
      Leaf of 'a
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    90
    | Node of ('a tree) list;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    91
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    92
  (* ('a -> 'b) -> 'a tree -> 'b tree *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    93
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    94
  fun tree_map f tr =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    95
    case tr of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    96
      Leaf x  => Leaf (f x)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    97
    | Node xs => Node (map (tree_map f) xs);
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    98
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    99
  (* ('a * 'b -> 'a) -> 'a * ('b tree) -> 'a *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   100
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   101
  fun tree_foldl f =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   102
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   103
    fun itl (e, Leaf x)  = f(e,x)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   104
      | itl (e, Node xs) = Library.foldl (tree_foldl f) (e,xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   105
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   106
    itl
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   107
  end;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   108
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   109
  (* 'a tree * 'b tree -> ('a * 'b) tree *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   110
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   111
  fun tree_pair (t1, t2) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   112
    case t1 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   113
      Leaf x =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   114
      (case t2 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   115
          Leaf y => Leaf (x,y)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   116
        | Node _ => raise REFUTE ("tree_pair",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   117
            "trees are of different height (second tree is higher)"))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   118
    | Node xs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   119
      (case t2 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   120
          (* '~~' will raise an exception if the number of branches in   *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   121
          (* both trees is different at the current node                 *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   122
          Node ys => Node (map tree_pair (xs ~~ ys))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   123
        | Leaf _  => raise REFUTE ("tree_pair",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   124
            "trees are of different height (first tree is higher)"));
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   125
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   126
(* ------------------------------------------------------------------------- *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   127
(* params: parameters that control the translation into a propositional      *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   128
(*         formula/model generation                                          *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   129
(*                                                                           *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   130
(* The following parameters are supported (and required (!), except for      *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   131
(* "sizes"):                                                                 *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   132
(*                                                                           *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   133
(* Name          Type    Description                                         *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   134
(*                                                                           *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   135
(* "sizes"       (string * int) list                                         *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   136
(*                       Size of ground types (e.g. 'a=2), or depth of IDTs. *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   137
(* "minsize"     int     If >0, minimal size of each ground type/IDT depth.  *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   138
(* "maxsize"     int     If >0, maximal size of each ground type/IDT depth.  *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   139
(* "maxvars"     int     If >0, use at most 'maxvars' Boolean variables      *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   140
(*                       when transforming the term into a propositional     *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   141
(*                       formula.                                            *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   142
(* "maxtime"     int     If >0, terminate after at most 'maxtime' seconds.   *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   143
(* "satsolver"   string  SAT solver to be used.                              *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   144
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   145
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   146
  type params =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   147
    {
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   148
      sizes    : (string * int) list,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   149
      minsize  : int,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   150
      maxsize  : int,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   151
      maxvars  : int,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   152
      maxtime  : int,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   153
      satsolver: string
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   154
    };
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   155
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   156
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   157
(* interpretation: a term's interpretation is given by a variable of type    *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   158
(*                 'interpretation'                                          *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   159
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   160
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   161
  type interpretation =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   162
    prop_formula list tree;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   163
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   164
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   165
(* model: a model specifies the size of types and the interpretation of      *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   166
(*        terms                                                              *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   167
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   168
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   169
  type model =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   170
    (Term.typ * int) list * (Term.term * interpretation) list;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   171
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   172
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   173
(* arguments: additional arguments required during interpretation of terms   *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   174
(* ------------------------------------------------------------------------- *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   175
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   176
  type arguments =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   177
    {
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   178
      (* just passed unchanged from 'params': *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   179
      maxvars   : int,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   180
      (* whether to use 'make_equality' or 'make_def_equality': *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   181
      def_eq    : bool,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   182
      (* the following may change during the translation: *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   183
      next_idx  : int,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   184
      bounds    : interpretation list,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   185
      wellformed: prop_formula
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   186
    };
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   187
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   188
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   189
  structure RefuteDataArgs =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   190
  struct
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   191
    val name = "HOL/refute";
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   192
    type T =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   193
      {interpreters: (string * (theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   194
        (interpretation * model * arguments) option)) list,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   195
       printers: (string * (theory -> model -> Term.term -> interpretation ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   196
        (int -> bool) -> Term.term option)) list,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   197
       parameters: string Symtab.table};
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   198
    val empty = {interpreters = [], printers = [], parameters = Symtab.empty};
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   199
    val copy = I;
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   200
    val extend = I;
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   201
    fun merge _
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   202
      ({interpreters = in1, printers = pr1, parameters = pa1},
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   203
       {interpreters = in2, printers = pr2, parameters = pa2}) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   204
      {interpreters = AList.merge (op =) (K true) (in1, in2),
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   205
       printers = AList.merge (op =) (K true) (pr1, pr2),
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   206
       parameters = Symtab.merge (op=) (pa1, pa2)};
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   207
    fun print sg {interpreters, printers, parameters} =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   208
      Pretty.writeln (Pretty.chunks
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   209
        [Pretty.strs ("default parameters:" :: List.concat (map
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   210
          (fn (name, value) => [name, "=", value]) (Symtab.dest parameters))),
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   211
         Pretty.strs ("interpreters:" :: map fst interpreters),
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   212
         Pretty.strs ("printers:" :: map fst printers)]);
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   213
  end;
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   214
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   215
  structure RefuteData = TheoryDataFun(RefuteDataArgs);
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   216
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   217
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   218
(* ------------------------------------------------------------------------- *)
15334
d5a92997dc1b exception CANNOT_INTERPRET removed (not needed anymore since the stlc_interpreter can interpret any term)
webertj
parents: 15333
diff changeset
   219
(* interpret: interprets the term 't' using a suitable interpreter; returns  *)
d5a92997dc1b exception CANNOT_INTERPRET removed (not needed anymore since the stlc_interpreter can interpret any term)
webertj
parents: 15333
diff changeset
   220
(*            the interpretation and a (possibly extended) model that keeps  *)
d5a92997dc1b exception CANNOT_INTERPRET removed (not needed anymore since the stlc_interpreter can interpret any term)
webertj
parents: 15333
diff changeset
   221
(*            track of the interpretation of subterms                        *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   222
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   223
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   224
  (* theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   225
    (interpretation * model * arguments) *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   226
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   227
  fun interpret thy model args t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   228
    case get_first (fn (_, f) => f thy model args t)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   229
      (#interpreters (RefuteData.get thy)) of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   230
      NONE   => raise REFUTE ("interpret",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   231
        "no interpreter for term " ^ quote (Sign.string_of_term thy t))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   232
    | SOME x => x;
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   233
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   234
(* ------------------------------------------------------------------------- *)
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
   235
(* print: converts the constant denoted by the term 't' into a term using a  *)
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
   236
(*        suitable printer                                                   *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   237
(* ------------------------------------------------------------------------- *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   238
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   239
  (* theory -> model -> Term.term -> interpretation -> (int -> bool) ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   240
    Term.term *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   241
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   242
  fun print thy model t intr assignment =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   243
    case get_first (fn (_, f) => f thy model t intr assignment)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   244
      (#printers (RefuteData.get thy)) of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   245
      NONE   => raise REFUTE ("print",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   246
        "no printer for term " ^ quote (Sign.string_of_term thy t))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   247
    | SOME x => x;
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   248
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   249
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   250
(* print_model: turns the model into a string, using a fixed interpretation  *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   251
(*              (given by an assignment for Boolean variables) and suitable  *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   252
(*              printers                                                     *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   253
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   254
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   255
  (* theory -> model -> (int -> bool) -> string *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   256
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   257
  fun print_model thy model assignment =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   258
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   259
    val (typs, terms) = model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   260
    val typs_msg =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   261
      if null typs then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   262
        "empty universe (no type variables in term)\n"
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   263
      else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   264
        "Size of types: " ^ commas (map (fn (T, i) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   265
          Sign.string_of_typ thy T ^ ": " ^ string_of_int i) typs) ^ "\n"
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   266
    val show_consts_msg =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   267
      if not (!show_consts) andalso Library.exists (is_Const o fst) terms then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   268
        "set \"show_consts\" to show the interpretation of constants\n"
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   269
      else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   270
        ""
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   271
    val terms_msg =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   272
      if null terms then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   273
        "empty interpretation (no free variables in term)\n"
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   274
      else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   275
        space_implode "\n" (List.mapPartial (fn (t, intr) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   276
          (* print constants only if 'show_consts' is true *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   277
          if (!show_consts) orelse not (is_Const t) then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   278
            SOME (Sign.string_of_term thy t ^ ": " ^
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   279
              Sign.string_of_term thy (print thy model t intr assignment))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   280
          else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   281
            NONE) terms) ^ "\n"
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   282
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   283
    typs_msg ^ show_consts_msg ^ terms_msg
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   284
  end;
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   285
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   286
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   287
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   288
(* PARAMETER MANAGEMENT                                                      *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   289
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   290
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   291
  (* string -> (theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   292
    (interpretation * model * arguments) option) -> theory -> theory *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   293
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   294
  fun add_interpreter name f thy =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   295
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   296
    val {interpreters, printers, parameters} = RefuteData.get thy
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   297
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   298
    case AList.lookup (op =) interpreters name of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   299
      NONE   => RefuteData.put {interpreters = (name, f) :: interpreters,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   300
      printers = printers, parameters = parameters} thy
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   301
    | SOME _ => error ("Interpreter " ^ name ^ " already declared")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   302
  end;
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   303
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   304
  (* string -> (theory -> model -> Term.term -> interpretation ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   305
    (int -> bool) -> Term.term option) -> theory -> theory *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   306
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   307
  fun add_printer name f thy =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   308
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   309
    val {interpreters, printers, parameters} = RefuteData.get thy
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   310
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   311
    case AList.lookup (op =) printers name of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   312
      NONE   => RefuteData.put {interpreters = interpreters,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   313
      printers = (name, f) :: printers, parameters = parameters} thy
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   314
    | SOME _ => error ("Printer " ^ name ^ " already declared")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   315
  end;
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   316
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   317
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   318
(* set_default_param: stores the '(name, value)' pair in RefuteData's        *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   319
(*                    parameter table                                        *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   320
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   321
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   322
  (* (string * string) -> theory -> theory *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   323
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   324
  fun set_default_param (name, value) thy =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   325
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   326
    val {interpreters, printers, parameters} = RefuteData.get thy
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   327
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   328
    RefuteData.put (case Symtab.lookup parameters name of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   329
      NONE   =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   330
      {interpreters = interpreters, printers = printers,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   331
        parameters = Symtab.extend (parameters, [(name, value)])}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   332
    | SOME _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   333
      {interpreters = interpreters, printers = printers,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   334
        parameters = Symtab.update (name, value) parameters}) thy
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   335
  end;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   336
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   337
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   338
(* get_default_param: retrieves the value associated with 'name' from        *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   339
(*                    RefuteData's parameter table                           *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   340
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   341
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   342
  (* theory -> string -> string option *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   343
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   344
  val get_default_param = Symtab.lookup o #parameters o RefuteData.get;
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   345
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   346
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   347
(* get_default_params: returns a list of all '(name, value)' pairs that are  *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   348
(*                     stored in RefuteData's parameter table                *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   349
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   350
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   351
  (* theory -> (string * string) list *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   352
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   353
  val get_default_params = Symtab.dest o #parameters o RefuteData.get;
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   354
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   355
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   356
(* actual_params: takes a (possibly empty) list 'params' of parameters that  *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   357
(*      override the default parameters currently specified in 'thy', and    *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   358
(*      returns a record that can be passed to 'find_model'.                 *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   359
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   360
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   361
  (* theory -> (string * string) list -> params *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   362
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   363
  fun actual_params thy override =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   364
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   365
    (* (string * string) list * string -> int *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   366
    fun read_int (parms, name) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   367
      case AList.lookup (op =) parms name of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   368
        SOME s => (case Int.fromString s of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   369
          SOME i => i
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   370
        | NONE   => error ("parameter " ^ quote name ^
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   371
          " (value is " ^ quote s ^ ") must be an integer value"))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   372
      | NONE   => error ("parameter " ^ quote name ^
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   373
          " must be assigned a value")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   374
    (* (string * string) list * string -> string *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   375
    fun read_string (parms, name) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   376
      case AList.lookup (op =) parms name of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   377
        SOME s => s
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   378
      | NONE   => error ("parameter " ^ quote name ^
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   379
        " must be assigned a value")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   380
    (* 'override' first, defaults last: *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   381
    (* (string * string) list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   382
    val allparams = override @ (get_default_params thy)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   383
    (* int *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   384
    val minsize   = read_int (allparams, "minsize")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   385
    val maxsize   = read_int (allparams, "maxsize")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   386
    val maxvars   = read_int (allparams, "maxvars")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   387
    val maxtime   = read_int (allparams, "maxtime")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   388
    (* string *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   389
    val satsolver = read_string (allparams, "satsolver")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   390
    (* all remaining parameters of the form "string=int" are collected in *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   391
    (* 'sizes'                                                            *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   392
    (* TODO: it is currently not possible to specify a size for a type    *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   393
    (*       whose name is one of the other parameters (e.g. 'maxvars')   *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   394
    (* (string * int) list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   395
    val sizes     = List.mapPartial
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   396
      (fn (name, value) => Option.map (pair name) (Int.fromString value))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   397
      (List.filter (fn (name, _) => name<>"minsize" andalso name<>"maxsize"
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   398
        andalso name<>"maxvars" andalso name<>"maxtime"
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   399
        andalso name<>"satsolver") allparams)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   400
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   401
    {sizes=sizes, minsize=minsize, maxsize=maxsize, maxvars=maxvars,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   402
      maxtime=maxtime, satsolver=satsolver}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   403
  end;
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   404
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   405
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   406
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   407
(* TRANSLATION HOL -> PROPOSITIONAL LOGIC, BOOLEAN ASSIGNMENT -> MODEL       *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   408
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   409
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   410
  (* (''a * 'b) list -> ''a -> 'b *)
22092
ab3dfcef6489 reformatted to 80 chars/line
webertj
parents: 22055
diff changeset
   411
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   412
  fun lookup xs key =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   413
    Option.valOf (AList.lookup (op =) xs key);
22092
ab3dfcef6489 reformatted to 80 chars/line
webertj
parents: 22055
diff changeset
   414
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   415
(* ------------------------------------------------------------------------- *)
15335
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
   416
(* typ_of_dtyp: converts a data type ('DatatypeAux.dtyp') into a type        *)
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
   417
(*              ('Term.typ'), given type parameters for the data type's type *)
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
   418
(*              arguments                                                    *)
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
   419
(* ------------------------------------------------------------------------- *)
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
   420
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   421
  (* DatatypeAux.descr -> (DatatypeAux.dtyp * Term.typ) list ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   422
    DatatypeAux.dtyp -> Term.typ *)
15335
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
   423
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   424
  fun typ_of_dtyp descr typ_assoc (DatatypeAux.DtTFree a) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   425
    (* replace a 'DtTFree' variable by the associated type *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   426
    lookup typ_assoc (DatatypeAux.DtTFree a)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   427
    | typ_of_dtyp descr typ_assoc (DatatypeAux.DtType (s, ds)) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   428
    Type (s, map (typ_of_dtyp descr typ_assoc) ds)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   429
    | typ_of_dtyp descr typ_assoc (DatatypeAux.DtRec i) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   430
    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   431
      val (s, ds, _) = lookup descr i
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   432
    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   433
      Type (s, map (typ_of_dtyp descr typ_assoc) ds)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   434
    end;
15335
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
   435
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
   436
(* ------------------------------------------------------------------------- *)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   437
(* close_form: universal closure over schematic variables in 't'             *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   438
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   439
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   440
  (* Term.term -> Term.term *)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   441
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   442
  fun close_form t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   443
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   444
    (* (Term.indexname * Term.typ) list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   445
    val vars = sort_wrt (fst o fst) (map dest_Var (term_vars t))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   446
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   447
    Library.foldl (fn (t', ((x, i), T)) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   448
      (Term.all T) $ Abs (x, T, abstract_over (Var ((x, i), T), t')))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   449
      (t, vars)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   450
  end;
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   451
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   452
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   453
(* monomorphic_term: applies a type substitution 'typeSubs' for all type     *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   454
(*                   variables in a term 't'                                 *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   455
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   456
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   457
  (* Type.tyenv -> Term.term -> Term.term *)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   458
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   459
  fun monomorphic_term typeSubs t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   460
    map_types (map_type_tvar
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   461
      (fn v =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   462
        case Type.lookup (typeSubs, v) of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   463
          NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   464
          (* schematic type variable not instantiated *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   465
          raise REFUTE ("monomorphic_term",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   466
            "no substitution for type variable " ^ fst (fst v) ^
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   467
            " in term " ^ Display.raw_string_of_term t)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   468
        | SOME typ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   469
          typ)) t;
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   470
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   471
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   472
(* specialize_type: given a constant 's' of type 'T', which is a subterm of  *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   473
(*                  't', where 't' has a (possibly) more general type, the   *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   474
(*                  schematic type variables in 't' are instantiated to      *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   475
(*                  match the type 'T' (may raise Type.TYPE_MATCH)           *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   476
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   477
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   478
  (* theory -> (string * Term.typ) -> Term.term -> Term.term *)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   479
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   480
  fun specialize_type thy (s, T) t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   481
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   482
    fun find_typeSubs (Const (s', T')) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   483
      if s=s' then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   484
        SOME (Sign.typ_match thy (T', T) Vartab.empty)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   485
          handle Type.TYPE_MATCH => NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   486
      else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   487
        NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   488
      | find_typeSubs (Free _)           = NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   489
      | find_typeSubs (Var _)            = NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   490
      | find_typeSubs (Bound _)          = NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   491
      | find_typeSubs (Abs (_, _, body)) = find_typeSubs body
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   492
      | find_typeSubs (t1 $ t2)          =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   493
      (case find_typeSubs t1 of SOME x => SOME x
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   494
                              | NONE   => find_typeSubs t2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   495
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   496
    case find_typeSubs t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   497
      SOME typeSubs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   498
      monomorphic_term typeSubs t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   499
    | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   500
      (* no match found - perhaps due to sort constraints *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   501
      raise Type.TYPE_MATCH
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   502
  end;
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   503
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   504
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   505
(* is_const_of_class: returns 'true' iff 'Const (s, T)' is a constant that   *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   506
(*                    denotes membership to an axiomatic type class          *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   507
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   508
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   509
  (* theory -> string * Term.typ -> bool *)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   510
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   511
  fun is_const_of_class thy (s, T) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   512
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   513
    val class_const_names = map Logic.const_of_class (Sign.all_classes thy)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   514
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   515
    (* I'm not quite sure if checking the name 's' is sufficient, *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   516
    (* or if we should also check the type 'T'.                   *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   517
    s mem_string class_const_names
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   518
  end;
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   519
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   520
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   521
(* is_IDT_constructor: returns 'true' iff 'Const (s, T)' is the constructor  *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   522
(*                     of an inductive datatype in 'thy'                     *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   523
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   524
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   525
  (* theory -> string * Term.typ -> bool *)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   526
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   527
  fun is_IDT_constructor thy (s, T) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   528
    (case body_type T of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   529
      Type (s', _) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   530
      (case DatatypePackage.get_datatype_constrs thy s' of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   531
        SOME constrs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   532
        List.exists (fn (cname, cty) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   533
          cname = s andalso Sign.typ_instance thy (T, cty)) constrs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   534
      | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   535
        false)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   536
    | _  =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   537
      false);
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   538
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   539
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   540
(* is_IDT_recursor: returns 'true' iff 'Const (s, T)' is the recursion       *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   541
(*                  operator of an inductive datatype in 'thy'               *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   542
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   543
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   544
  (* theory -> string * Term.typ -> bool *)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   545
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   546
  fun is_IDT_recursor thy (s, T) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   547
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   548
    val rec_names = Symtab.fold (append o #rec_names o snd)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   549
      (DatatypePackage.get_datatypes thy) []
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   550
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   551
    (* I'm not quite sure if checking the name 's' is sufficient, *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   552
    (* or if we should also check the type 'T'.                   *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   553
    s mem_string rec_names
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   554
  end;
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   555
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   556
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   557
(* get_def: looks up the definition of a constant, as created by "constdefs" *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   558
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   559
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   560
  (* theory -> string * Term.typ -> (string * Term.term) option *)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   561
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   562
  fun get_def thy (s, T) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   563
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   564
    (* maps  f ?t1 ... ?tn == rhs  to  %t1...tn. rhs *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   565
    fun norm_rhs eqn =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   566
    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   567
      fun lambda (v as Var ((x, _), T)) t = Abs (x, T, abstract_over (v, t))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   568
        | lambda v t                      = raise TERM ("lambda", [v, t])
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   569
      val (lhs, rhs) = Logic.dest_equals eqn
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   570
      val (_, args)  = Term.strip_comb lhs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   571
    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   572
      fold lambda (rev args) rhs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   573
    end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   574
    (* (string * Term.term) list -> (string * Term.term) option *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   575
    fun get_def_ax [] = NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   576
      | get_def_ax ((axname, ax) :: axioms) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   577
      (let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   578
        val (lhs, _) = Logic.dest_equals ax  (* equations only *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   579
        val c        = Term.head_of lhs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   580
        val (s', T') = Term.dest_Const c
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   581
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   582
        if s=s' then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   583
          let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   584
            val typeSubs = Sign.typ_match thy (T', T) Vartab.empty
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   585
            val ax'      = monomorphic_term typeSubs ax
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   586
            val rhs      = norm_rhs ax'
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   587
          in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   588
            SOME (axname, rhs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   589
          end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   590
        else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   591
          get_def_ax axioms
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   592
      end handle ERROR _         => get_def_ax axioms
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   593
               | TERM _          => get_def_ax axioms
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   594
               | Type.TYPE_MATCH => get_def_ax axioms)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   595
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   596
    get_def_ax (Theory.all_axioms_of thy)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   597
  end;
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   598
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   599
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   600
(* get_typedef: looks up the definition of a type, as created by "typedef"   *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   601
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   602
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   603
  (* theory -> (string * Term.typ) -> (string * Term.term) option *)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   604
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   605
  fun get_typedef thy T =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   606
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   607
    (* (string * Term.term) list -> (string * Term.term) option *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   608
    fun get_typedef_ax [] = NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   609
      | get_typedef_ax ((axname, ax) :: axioms) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   610
      (let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   611
        (* Term.term -> Term.typ option *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   612
        fun type_of_type_definition (Const (s', T')) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   613
          if s'="Typedef.type_definition" then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   614
            SOME T'
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   615
          else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   616
            NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   617
          | type_of_type_definition (Free _)           = NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   618
          | type_of_type_definition (Var _)            = NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   619
          | type_of_type_definition (Bound _)          = NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   620
          | type_of_type_definition (Abs (_, _, body)) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   621
          type_of_type_definition body
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   622
          | type_of_type_definition (t1 $ t2)          =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   623
          (case type_of_type_definition t1 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   624
            SOME x => SOME x
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   625
          | NONE   => type_of_type_definition t2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   626
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   627
        case type_of_type_definition ax of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   628
          SOME T' =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   629
          let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   630
            val T''      = (domain_type o domain_type) T'
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   631
            val typeSubs = Sign.typ_match thy (T'', T) Vartab.empty
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   632
          in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   633
            SOME (axname, monomorphic_term typeSubs ax)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   634
          end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   635
        | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   636
          get_typedef_ax axioms
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   637
      end handle ERROR _         => get_typedef_ax axioms
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   638
               | MATCH           => get_typedef_ax axioms
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   639
               | Type.TYPE_MATCH => get_typedef_ax axioms)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   640
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   641
    get_typedef_ax (Theory.all_axioms_of thy)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   642
  end;
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   643
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   644
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   645
(* get_classdef: looks up the defining axiom for an axiomatic type class, as *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   646
(*               created by the "axclass" command                            *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   647
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   648
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   649
  (* theory -> string -> (string * Term.term) option *)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   650
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   651
  fun get_classdef thy class =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   652
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   653
    val axname = class ^ "_class_def"
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   654
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   655
    Option.map (pair axname)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   656
      (AList.lookup (op =) (Theory.all_axioms_of thy) axname)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   657
  end;
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   658
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   659
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   660
(* unfold_defs: unfolds all defined constants in a term 't', beta-eta        *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   661
(*              normalizes the result term; certain constants are not        *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   662
(*              unfolded (cf. 'collect_axioms' and the various interpreters  *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   663
(*              below): if the interpretation respects a definition anyway,  *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   664
(*              that definition does not need to be unfolded                 *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   665
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   666
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   667
  (* theory -> Term.term -> Term.term *)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   668
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   669
  (* Note: we could intertwine unfolding of constants and beta-(eta-)       *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   670
  (*       normalization; this would save some unfolding for terms where    *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   671
  (*       constants are eliminated by beta-reduction (e.g. 'K c1 c2').  On *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   672
  (*       the other hand, this would cause additional work for terms where *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   673
  (*       constants are duplicated by beta-reduction (e.g. 'S c1 c2 c3').  *)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   674
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   675
  fun unfold_defs thy t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   676
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   677
    (* Term.term -> Term.term *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   678
    fun unfold_loop t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   679
      case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   680
      (* Pure *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   681
        Const ("all", _)                => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   682
      | Const ("==", _)                 => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   683
      | Const ("==>", _)                => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   684
      | Const ("TYPE", _)               => t  (* axiomatic type classes *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   685
      (* HOL *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   686
      | Const ("Trueprop", _)           => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   687
      | Const ("Not", _)                => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   688
      | (* redundant, since 'True' is also an IDT constructor *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   689
        Const ("True", _)               => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   690
      | (* redundant, since 'False' is also an IDT constructor *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   691
        Const ("False", _)              => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   692
      | Const ("arbitrary", _)          => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   693
      | Const ("The", _)                => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   694
      | Const ("Hilbert_Choice.Eps", _) => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   695
      | Const ("All", _)                => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   696
      | Const ("Ex", _)                 => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   697
      | Const ("op =", _)               => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   698
      | Const ("op &", _)               => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   699
      | Const ("op |", _)               => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   700
      | Const ("op -->", _)             => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   701
      (* sets *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   702
      | Const ("Collect", _)            => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   703
      | Const ("op :", _)               => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   704
      (* other optimizations *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   705
      | Const ("Finite_Set.card", _)    => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   706
      | Const ("Finite_Set.Finites", _) => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   707
      | Const ("Finite_Set.finite", _)  => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   708
      | Const ("Orderings.less", Type ("fun", [Type ("nat", []),
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   709
        Type ("fun", [Type ("nat", []), Type ("bool", [])])])) => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   710
      | Const ("HOL.plus", Type ("fun", [Type ("nat", []),
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   711
        Type ("fun", [Type ("nat", []), Type ("nat", [])])])) => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   712
      | Const ("HOL.minus", Type ("fun", [Type ("nat", []),
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   713
        Type ("fun", [Type ("nat", []), Type ("nat", [])])])) => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   714
      | Const ("HOL.times", Type ("fun", [Type ("nat", []),
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   715
        Type ("fun", [Type ("nat", []), Type ("nat", [])])])) => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   716
      | Const ("List.op @", _)          => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   717
      | Const ("Lfp.lfp", _)            => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   718
      | Const ("Gfp.gfp", _)            => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   719
      | Const ("fst", _)                => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   720
      | Const ("snd", _)                => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   721
      (* simply-typed lambda calculus *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   722
      | Const (s, T) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   723
        (if is_IDT_constructor thy (s, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   724
          orelse is_IDT_recursor thy (s, T) then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   725
          t  (* do not unfold IDT constructors/recursors *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   726
        (* unfold the constant if there is a defining equation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   727
        else case get_def thy (s, T) of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   728
          SOME (axname, rhs) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   729
          (* Note: if the term to be unfolded (i.e. 'Const (s, T)')  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   730
          (* occurs on the right-hand side of the equation, i.e. in  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   731
          (* 'rhs', we must not use this equation to unfold, because *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   732
          (* that would loop.  Here would be the right place to      *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   733
          (* check this.  However, getting this really right seems   *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   734
          (* difficult because the user may state arbitrary axioms,  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   735
          (* which could interact with overloading to create loops.  *)
22580
d91b4dd651d6 cleaned-up Output functions;
wenzelm
parents: 22567
diff changeset
   736
          ((*Output.immediate_output (" unfolding: " ^ axname);*)unfold_loop rhs)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   737
        | NONE => t)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   738
      | Free _           => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   739
      | Var _            => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   740
      | Bound _          => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   741
      | Abs (s, T, body) => Abs (s, T, unfold_loop body)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   742
      | t1 $ t2          => (unfold_loop t1) $ (unfold_loop t2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   743
    val result = Envir.beta_eta_contract (unfold_loop t)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   744
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   745
    result
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   746
  end;
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   747
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   748
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   749
(* collect_axioms: collects (monomorphic, universally quantified, unfolded   *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   750
(*                 versions of) all HOL axioms that are relevant w.r.t 't'   *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   751
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   752
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   753
  (* Note: to make the collection of axioms more easily extensible, this    *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   754
  (*       function could be based on user-supplied "axiom collectors",     *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   755
  (*       similar to 'interpret'/interpreters or 'print'/printers          *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   756
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   757
  (* Note: currently we use "inverse" functions to the definitional         *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   758
  (*       mechanisms provided by Isabelle/HOL, e.g. for "axclass",         *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   759
  (*       "typedef", "constdefs".  A more general approach could consider  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   760
  (*       *every* axiom of the theory and collect it if it has a constant/ *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   761
  (*       type/typeclass in common with the term 't'.                      *)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   762
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   763
  (* theory -> Term.term -> Term.term list *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   764
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   765
  (* Which axioms are "relevant" for a particular term/type goes hand in    *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   766
  (* hand with the interpretation of that term/type by its interpreter (see *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   767
  (* way below): if the interpretation respects an axiom anyway, the axiom  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   768
  (* does not need to be added as a constraint here.                        *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   769
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   770
  (* To avoid collecting the same axiom multiple times, we use an           *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   771
  (* accumulator 'axs' which contains all axioms collected so far.          *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   772
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   773
  fun collect_axioms thy t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   774
  let
22580
d91b4dd651d6 cleaned-up Output functions;
wenzelm
parents: 22567
diff changeset
   775
    val _ = Output.immediate_output "Adding axioms..."
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   776
    (* (string * Term.term) list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   777
    val axioms = Theory.all_axioms_of thy
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   778
    (* string * Term.term -> Term.term list -> Term.term list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   779
    fun collect_this_axiom (axname, ax) axs =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   780
    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   781
      val ax' = unfold_defs thy ax
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   782
    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   783
      if member (op aconv) axs ax' then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   784
        axs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   785
      else (
22580
d91b4dd651d6 cleaned-up Output functions;
wenzelm
parents: 22567
diff changeset
   786
        Output.immediate_output (" " ^ axname);
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   787
        collect_term_axioms (ax' :: axs, ax')
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   788
      )
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   789
    end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   790
    (* Term.term list * Term.typ -> Term.term list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   791
    and collect_sort_axioms (axs, T) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   792
    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   793
      (* string list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   794
      val sort = (case T of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   795
          TFree (_, sort) => sort
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   796
        | TVar (_, sort)  => sort
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   797
        | _               => raise REFUTE ("collect_axioms", "type " ^
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   798
          Sign.string_of_typ thy T ^ " is not a variable"))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   799
      (* obtain axioms for all superclasses *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   800
      val superclasses = sort @ (maps (Sign.super_classes thy) sort)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   801
      (* merely an optimization, because 'collect_this_axiom' disallows *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   802
      (* duplicate axioms anyway:                                       *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   803
      val superclasses = distinct (op =) superclasses
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   804
      val class_axioms = maps (fn class => map (fn ax =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   805
        ("<" ^ class ^ ">", Thm.prop_of ax))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   806
        (#axioms (AxClass.get_definition thy class) handle ERROR _ => []))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   807
        superclasses
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   808
      (* replace the (at most one) schematic type variable in each axiom *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   809
      (* by the actual type 'T'                                          *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   810
      val monomorphic_class_axioms = map (fn (axname, ax) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   811
        (case Term.term_tvars ax of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   812
          [] =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   813
          (axname, ax)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   814
        | [(idx, S)] =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   815
          (axname, monomorphic_term (Vartab.make [(idx, (S, T))]) ax)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   816
        | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   817
          raise REFUTE ("collect_axioms", "class axiom " ^ axname ^ " (" ^
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   818
            Sign.string_of_term thy ax ^
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   819
            ") contains more than one type variable")))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   820
        class_axioms
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   821
    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   822
      fold collect_this_axiom monomorphic_class_axioms axs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   823
    end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   824
    (* Term.term list * Term.typ -> Term.term list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   825
    and collect_type_axioms (axs, T) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   826
      case T of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   827
      (* simple types *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   828
        Type ("prop", [])      => axs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   829
      | Type ("fun", [T1, T2]) => collect_type_axioms
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   830
        (collect_type_axioms (axs, T1), T2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   831
      | Type ("set", [T1])     => collect_type_axioms (axs, T1)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   832
      (* axiomatic type classes *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   833
      | Type ("itself", [T1])  => collect_type_axioms (axs, T1)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   834
      | Type (s, Ts)           =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   835
        (case DatatypePackage.get_datatype thy s of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   836
          SOME info =>  (* inductive datatype *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   837
            (* only collect relevant type axioms for the argument types *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   838
            Library.foldl collect_type_axioms (axs, Ts)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   839
        | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   840
          (case get_typedef thy T of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   841
            SOME (axname, ax) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   842
            collect_this_axiom (axname, ax) axs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   843
          | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   844
            (* unspecified type, perhaps introduced with "typedecl" *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   845
            (* at least collect relevant type axioms for the argument types *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   846
            Library.foldl collect_type_axioms (axs, Ts)))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   847
      (* axiomatic type classes *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   848
      | TFree _                => collect_sort_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   849
      (* axiomatic type classes *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   850
      | TVar _                 => collect_sort_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   851
    (* Term.term list * Term.term -> Term.term list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   852
    and collect_term_axioms (axs, t) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   853
      case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   854
      (* Pure *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   855
        Const ("all", _)                => axs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   856
      | Const ("==", _)                 => axs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   857
      | Const ("==>", _)                => axs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   858
      (* axiomatic type classes *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   859
      | Const ("TYPE", T)               => collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   860
      (* HOL *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   861
      | Const ("Trueprop", _)           => axs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   862
      | Const ("Not", _)                => axs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   863
      (* redundant, since 'True' is also an IDT constructor *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   864
      | Const ("True", _)               => axs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   865
      (* redundant, since 'False' is also an IDT constructor *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   866
      | Const ("False", _)              => axs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   867
      | Const ("arbitrary", T)          => collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   868
      | Const ("The", T)                =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   869
        let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   870
          val ax = specialize_type thy ("The", T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   871
            (lookup axioms "HOL.the_eq_trivial")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   872
        in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   873
          collect_this_axiom ("HOL.the_eq_trivial", ax) axs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   874
        end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   875
      | Const ("Hilbert_Choice.Eps", T) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   876
        let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   877
          val ax = specialize_type thy ("Hilbert_Choice.Eps", T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   878
            (lookup axioms "Hilbert_Choice.someI")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   879
        in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   880
          collect_this_axiom ("Hilbert_Choice.someI", ax) axs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   881
        end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   882
      | Const ("All", T)                => collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   883
      | Const ("Ex", T)                 => collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   884
      | Const ("op =", T)               => collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   885
      | Const ("op &", _)               => axs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   886
      | Const ("op |", _)               => axs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   887
      | Const ("op -->", _)             => axs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   888
      (* sets *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   889
      | Const ("Collect", T)            => collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   890
      | Const ("op :", T)               => collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   891
      (* other optimizations *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   892
      | Const ("Finite_Set.card", T)    => collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   893
      | Const ("Finite_Set.Finites", T) => collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   894
      | Const ("Finite_Set.finite", T)  => collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   895
      | Const ("Orderings.less", T as Type ("fun", [Type ("nat", []),
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   896
        Type ("fun", [Type ("nat", []), Type ("bool", [])])])) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   897
          collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   898
      | Const ("HOL.plus", T as Type ("fun", [Type ("nat", []),
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   899
        Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   900
          collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   901
      | Const ("HOL.minus", T as Type ("fun", [Type ("nat", []),
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   902
        Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   903
          collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   904
      | Const ("HOL.times", T as Type ("fun", [Type ("nat", []),
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   905
        Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   906
          collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   907
      | Const ("List.op @", T)          => collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   908
      | Const ("Lfp.lfp", T)            => collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   909
      | Const ("Gfp.gfp", T)            => collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   910
      | Const ("fst", T)                => collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   911
      | Const ("snd", T)                => collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   912
      (* simply-typed lambda calculus *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   913
      | Const (s, T)                    =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   914
          if is_const_of_class thy (s, T) then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   915
            (* axiomatic type classes: add "OFCLASS(?'a::c, c_class)" *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   916
            (* and the class definition                               *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   917
            let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   918
              val class   = Logic.class_of_const s
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   919
              val inclass = Logic.mk_inclass (TVar (("'a", 0), [class]), class)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   920
              val ax_in   = SOME (specialize_type thy (s, T) inclass)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   921
                (* type match may fail due to sort constraints *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   922
                handle Type.TYPE_MATCH => NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   923
              val ax_1 = Option.map (fn ax => (Sign.string_of_term thy ax, ax))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   924
                ax_in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   925
              val ax_2 = Option.map (apsnd (specialize_type thy (s, T)))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   926
                (get_classdef thy class)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   927
            in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   928
              collect_type_axioms (fold collect_this_axiom
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   929
                (map_filter I [ax_1, ax_2]) axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   930
            end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   931
          else if is_IDT_constructor thy (s, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   932
            orelse is_IDT_recursor thy (s, T) then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   933
            (* only collect relevant type axioms *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   934
            collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   935
          else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   936
            (* other constants should have been unfolded, with some *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   937
            (* exceptions: e.g. Abs_xxx/Rep_xxx functions for       *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   938
            (* typedefs, or type-class related constants            *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   939
            (* only collect relevant type axioms *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   940
            collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   941
      | Free (_, T)      => collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   942
      | Var (_, T)       => collect_type_axioms (axs, T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   943
      | Bound i          => axs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   944
      | Abs (_, T, body) => collect_term_axioms
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   945
        (collect_type_axioms (axs, T), body)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   946
      | t1 $ t2          => collect_term_axioms
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   947
        (collect_term_axioms (axs, t1), t2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   948
    (* Term.term list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   949
    val result = map close_form (collect_term_axioms ([], t))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   950
    val _ = writeln " ...done."
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   951
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   952
    result
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   953
  end;
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   954
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   955
(* ------------------------------------------------------------------------- *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   956
(* ground_types: collects all ground types in a term (including argument     *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   957
(*               types of other types), suppressing duplicates.  Does not    *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   958
(*               return function types, set types, non-recursive IDTs, or    *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   959
(*               'propT'.  For IDTs, also the argument types of constructors *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   960
(*               are considered.                                             *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   961
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   962
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   963
  (* theory -> Term.term -> Term.typ list *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
   964
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   965
  fun ground_types thy t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   966
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   967
    (* Term.typ * Term.typ list -> Term.typ list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   968
    fun collect_types (T, acc) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   969
      if T mem acc then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   970
        acc  (* prevent infinite recursion (for IDTs) *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   971
      else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   972
        (case T of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   973
          Type ("fun", [T1, T2]) => collect_types (T1, collect_types (T2, acc))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   974
        | Type ("prop", [])      => acc
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   975
        | Type ("set", [T1])     => collect_types (T1, acc)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   976
        | Type (s, Ts)           =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   977
          (case DatatypePackage.get_datatype thy s of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   978
            SOME info =>  (* inductive datatype *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   979
            let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   980
              val index               = #index info
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   981
              val descr               = #descr info
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   982
              val (_, dtyps, constrs) = lookup descr index
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   983
              val typ_assoc           = dtyps ~~ Ts
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   984
              (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   985
              val _ = (if Library.exists (fn d =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   986
                  case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   987
                then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   988
                  raise REFUTE ("ground_types", "datatype argument (for type "
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   989
                    ^ Sign.string_of_typ thy (Type (s, Ts))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   990
                    ^ ") is not a variable")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   991
                else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   992
                  ())
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   993
              (* if the current type is a recursive IDT (i.e. a depth is *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   994
              (* required), add it to 'acc'                              *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   995
              val acc' = (if Library.exists (fn (_, ds) => Library.exists
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   996
                DatatypeAux.is_rec_type ds) constrs then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   997
                  insert (op =) T acc
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   998
                else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   999
                  acc)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1000
              (* collect argument types *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1001
              val acc_args = foldr collect_types acc' Ts
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1002
              (* collect constructor types *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1003
              val acc_constrs = foldr collect_types acc_args (List.concat
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1004
                (map (fn (_, ds) => map (typ_of_dtyp descr typ_assoc) ds)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1005
                  constrs))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1006
            in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1007
              acc_constrs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1008
            end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1009
          | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1010
            (* not an inductive datatype, e.g. defined via "typedef" or *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1011
            (* "typedecl"                                               *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1012
            insert (op =) T (foldr collect_types acc Ts))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1013
        | TFree _                => insert (op =) T acc
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1014
        | TVar _                 => insert (op =) T acc)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1015
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1016
    it_term_types collect_types (t, [])
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1017
  end;
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1018
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1019
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1020
(* string_of_typ: (rather naive) conversion from types to strings, used to   *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1021
(*                look up the size of a type in 'sizes'.  Parameterized      *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1022
(*                types with different parameters (e.g. "'a list" vs. "bool  *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1023
(*                list") are identified.                                     *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1024
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1025
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1026
  (* Term.typ -> string *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1027
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1028
  fun string_of_typ (Type (s, _))     = s
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1029
    | string_of_typ (TFree (s, _))    = s
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1030
    | string_of_typ (TVar ((s,_), _)) = s;
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1031
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1032
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1033
(* first_universe: returns the "first" (i.e. smallest) universe by assigning *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1034
(*                 'minsize' to every type for which no size is specified in *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1035
(*                 'sizes'                                                   *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1036
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1037
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1038
  (* Term.typ list -> (string * int) list -> int -> (Term.typ * int) list *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1039
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1040
  fun first_universe xs sizes minsize =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1041
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1042
    fun size_of_typ T =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1043
      case AList.lookup (op =) sizes (string_of_typ T) of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1044
        SOME n => n
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1045
      | NONE   => minsize
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1046
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1047
    map (fn T => (T, size_of_typ T)) xs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1048
  end;
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1049
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1050
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1051
(* next_universe: enumerates all universes (i.e. assignments of sizes to     *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1052
(*                types), where the minimal size of a type is given by       *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1053
(*                'minsize', the maximal size is given by 'maxsize', and a   *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1054
(*                type may have a fixed size given in 'sizes'                *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1055
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1056
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1057
  (* (Term.typ * int) list -> (string * int) list -> int -> int ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1058
    (Term.typ * int) list option *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1059
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1060
  fun next_universe xs sizes minsize maxsize =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1061
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1062
    (* creates the "first" list of length 'len', where the sum of all list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1063
    (* elements is 'sum', and the length of the list is 'len'              *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1064
    (* int -> int -> int -> int list option *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1065
    fun make_first _ 0 sum =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1066
      if sum=0 then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1067
        SOME []
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1068
      else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1069
        NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1070
      | make_first max len sum =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1071
      if sum<=max orelse max<0 then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1072
        Option.map (fn xs' => sum :: xs') (make_first max (len-1) 0)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1073
      else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1074
        Option.map (fn xs' => max :: xs') (make_first max (len-1) (sum-max))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1075
    (* enumerates all int lists with a fixed length, where 0<=x<='max' for *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1076
    (* all list elements x (unless 'max'<0)                                *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1077
    (* int -> int -> int -> int list -> int list option *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1078
    fun next max len sum [] =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1079
      NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1080
      | next max len sum [x] =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1081
      (* we've reached the last list element, so there's no shift possible *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1082
      make_first max (len+1) (sum+x+1)  (* increment 'sum' by 1 *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1083
      | next max len sum (x1::x2::xs) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1084
      if x1>0 andalso (x2<max orelse max<0) then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1085
        (* we can shift *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1086
        SOME (valOf (make_first max (len+1) (sum+x1-1)) @ (x2+1) :: xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1087
      else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1088
        (* continue search *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1089
        next max (len+1) (sum+x1) (x2::xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1090
    (* only consider those types for which the size is not fixed *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1091
    val mutables = List.filter
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1092
      (not o (AList.defined (op =) sizes) o string_of_typ o fst) xs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1093
    (* subtract 'minsize' from every size (will be added again at the end) *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1094
    val diffs = map (fn (_, n) => n-minsize) mutables
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1095
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1096
    case next (maxsize-minsize) 0 0 diffs of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1097
      SOME diffs' =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1098
      (* merge with those types for which the size is fixed *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1099
      SOME (snd (foldl_map (fn (ds, (T, _)) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1100
        case AList.lookup (op =) sizes (string_of_typ T) of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1101
        (* return the fixed size *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1102
          SOME n => (ds, (T, n))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1103
        (* consume the head of 'ds', add 'minsize' *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1104
        | NONE   => (tl ds, (T, minsize + hd ds)))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1105
        (diffs', xs)))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1106
    | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1107
      NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1108
  end;
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1109
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1110
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1111
(* toTrue: converts the interpretation of a Boolean value to a propositional *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1112
(*         formula that is true iff the interpretation denotes "true"        *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1113
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1114
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1115
  (* interpretation -> prop_formula *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1116
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1117
  fun toTrue (Leaf [fm, _]) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1118
    fm
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1119
    | toTrue _              =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1120
    raise REFUTE ("toTrue", "interpretation does not denote a Boolean value");
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1121
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1122
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1123
(* toFalse: converts the interpretation of a Boolean value to a              *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1124
(*          propositional formula that is true iff the interpretation        *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1125
(*          denotes "false"                                                  *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1126
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1127
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1128
  (* interpretation -> prop_formula *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1129
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1130
  fun toFalse (Leaf [_, fm]) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1131
    fm
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1132
    | toFalse _              =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1133
    raise REFUTE ("toFalse", "interpretation does not denote a Boolean value");
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1134
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1135
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1136
(* find_model: repeatedly calls 'interpret' with appropriate parameters,     *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1137
(*             applies a SAT solver, and (in case a model is found) displays *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1138
(*             the model to the user by calling 'print_model'                *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1139
(* thy       : the current theory                                            *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1140
(* {...}     : parameters that control the translation/model generation      *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1141
(* t         : term to be translated into a propositional formula            *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1142
(* negate    : if true, find a model that makes 't' false (rather than true) *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1143
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1144
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1145
  (* theory -> params -> Term.term -> bool -> unit *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1146
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1147
  fun find_model thy {sizes, minsize, maxsize, maxvars, maxtime, satsolver} t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1148
    negate =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1149
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1150
    (* unit -> unit *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1151
    fun wrapper () =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1152
    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1153
      val u      = unfold_defs thy t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1154
      val _      = writeln ("Unfolded term: " ^ Sign.string_of_term thy u)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1155
      val axioms = collect_axioms thy u
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1156
      (* Term.typ list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1157
      val types = Library.foldl (fn (acc, t') =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1158
        acc union (ground_types thy t')) ([], u :: axioms)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1159
      val _     = writeln ("Ground types: "
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1160
        ^ (if null types then "none."
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1161
           else commas (map (Sign.string_of_typ thy) types)))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1162
      (* we can only consider fragments of recursive IDTs, so we issue a  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1163
      (* warning if the formula contains a recursive IDT                  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1164
      (* TODO: no warning needed for /positive/ occurrences of IDTs       *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1165
      val _ = if Library.exists (fn
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1166
          Type (s, _) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1167
          (case DatatypePackage.get_datatype thy s of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1168
            SOME info =>  (* inductive datatype *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1169
            let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1170
              val index           = #index info
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1171
              val descr           = #descr info
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1172
              val (_, _, constrs) = lookup descr index
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1173
            in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1174
              (* recursive datatype? *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1175
              Library.exists (fn (_, ds) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1176
                Library.exists DatatypeAux.is_rec_type ds) constrs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1177
            end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1178
          | NONE => false)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1179
        | _ => false) types then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1180
          warning ("Term contains a recursive datatype; "
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1181
            ^ "countermodel(s) may be spurious!")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1182
        else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1183
          ()
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1184
      (* (Term.typ * int) list -> unit *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1185
      fun find_model_loop universe =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1186
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1187
        val init_model = (universe, [])
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1188
        val init_args  = {maxvars = maxvars, def_eq = false, next_idx = 1,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1189
          bounds = [], wellformed = True}
22580
d91b4dd651d6 cleaned-up Output functions;
wenzelm
parents: 22567
diff changeset
  1190
        val _          = Output.immediate_output ("Translating term (sizes: "
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1191
          ^ commas (map (fn (_, n) => string_of_int n) universe) ^ ") ...")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1192
        (* translate 'u' and all axioms *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1193
        val ((model, args), intrs) = foldl_map (fn ((m, a), t') =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1194
          let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1195
            val (i, m', a') = interpret thy m a t'
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1196
          in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1197
            (* set 'def_eq' to 'true' *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1198
            ((m', {maxvars = #maxvars a', def_eq = true,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1199
              next_idx = #next_idx a', bounds = #bounds a',
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1200
              wellformed = #wellformed a'}), i)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1201
          end) ((init_model, init_args), u :: axioms)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1202
        (* make 'u' either true or false, and make all axioms true, and *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1203
        (* add the well-formedness side condition                       *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1204
        val fm_u  = (if negate then toFalse else toTrue) (hd intrs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1205
        val fm_ax = PropLogic.all (map toTrue (tl intrs))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1206
        val fm    = PropLogic.all [#wellformed args, fm_ax, fm_u]
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1207
      in
22580
d91b4dd651d6 cleaned-up Output functions;
wenzelm
parents: 22567
diff changeset
  1208
        Output.immediate_output " invoking SAT solver...";
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1209
        (case SatSolver.invoke_solver satsolver fm of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1210
          SatSolver.SATISFIABLE assignment =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1211
          (writeln " model found!";
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1212
          writeln ("*** Model found: ***\n" ^ print_model thy model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1213
            (fn i => case assignment i of SOME b => b | NONE => true)))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1214
        | SatSolver.UNSATISFIABLE _ =>
22580
d91b4dd651d6 cleaned-up Output functions;
wenzelm
parents: 22567
diff changeset
  1215
          (Output.immediate_output " no model exists.\n";
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1216
          case next_universe universe sizes minsize maxsize of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1217
            SOME universe' => find_model_loop universe'
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1218
          | NONE           => writeln
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1219
            "Search terminated, no larger universe within the given limits.")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1220
        | SatSolver.UNKNOWN =>
22580
d91b4dd651d6 cleaned-up Output functions;
wenzelm
parents: 22567
diff changeset
  1221
          (Output.immediate_output " no model found.\n";
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1222
          case next_universe universe sizes minsize maxsize of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1223
            SOME universe' => find_model_loop universe'
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1224
          | NONE           => writeln
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1225
            "Search terminated, no larger universe within the given limits.")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1226
        ) handle SatSolver.NOT_CONFIGURED =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1227
          error ("SAT solver " ^ quote satsolver ^ " is not configured.")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1228
      end handle MAXVARS_EXCEEDED =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1229
        writeln ("\nSearch terminated, number of Boolean variables ("
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1230
          ^ string_of_int maxvars ^ " allowed) exceeded.")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1231
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1232
        find_model_loop (first_universe types sizes minsize)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1233
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1234
    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1235
      (* some parameter sanity checks *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1236
      minsize>=1 orelse
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1237
        error ("\"minsize\" is " ^ string_of_int minsize ^ ", must be at least 1");
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1238
      maxsize>=1 orelse
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1239
        error ("\"maxsize\" is " ^ string_of_int maxsize ^ ", must be at least 1");
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1240
      maxsize>=minsize orelse
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1241
        error ("\"maxsize\" (=" ^ string_of_int maxsize ^
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1242
        ") is less than \"minsize\" (=" ^ string_of_int minsize ^ ").");
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1243
      maxvars>=0 orelse
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1244
        error ("\"maxvars\" is " ^ string_of_int maxvars ^ ", must be at least 0");
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1245
      maxtime>=0 orelse
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1246
        error ("\"maxtime\" is " ^ string_of_int maxtime ^ ", must be at least 0");
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1247
      (* enter loop with or without time limit *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1248
      writeln ("Trying to find a model that "
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1249
        ^ (if negate then "refutes" else "satisfies") ^ ": "
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1250
        ^ Sign.string_of_term thy t);
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1251
      if maxtime>0 then (
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1252
        interrupt_timeout (Time.fromSeconds (Int.toLarge maxtime))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1253
          wrapper ()
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1254
        handle Interrupt =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1255
          writeln ("\nSearch terminated, time limit (" ^ string_of_int maxtime
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1256
            ^ (if maxtime=1 then " second" else " seconds") ^ ") exceeded.")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1257
      ) else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1258
        wrapper ()
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1259
    end;
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1260
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1261
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1262
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1263
(* INTERFACE, PART 2: FINDING A MODEL                                        *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1264
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1265
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1266
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1267
(* satisfy_term: calls 'find_model' to find a model that satisfies 't'       *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1268
(* params      : list of '(name, value)' pairs used to override default      *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1269
(*               parameters                                                  *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1270
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1271
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1272
  (* theory -> (string * string) list -> Term.term -> unit *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1273
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1274
  fun satisfy_term thy params t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1275
    find_model thy (actual_params thy params) t false;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1276
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1277
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1278
(* refute_term: calls 'find_model' to find a model that refutes 't'          *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1279
(* params     : list of '(name, value)' pairs used to override default       *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1280
(*              parameters                                                   *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1281
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1282
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1283
  (* theory -> (string * string) list -> Term.term -> unit *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1284
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1285
  fun refute_term thy params t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1286
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1287
    (* disallow schematic type variables, since we cannot properly negate  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1288
    (* terms containing them (their logical meaning is that there EXISTS a *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1289
    (* type s.t. ...; to refute such a formula, we would have to show that *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1290
    (* for ALL types, not ...)                                             *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1291
    val _ = null (term_tvars t) orelse
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1292
      error "Term to be refuted contains schematic type variables"
21556
e0ffb2d13f9f outermost universal quantifiers are stripped
webertj
parents: 21267
diff changeset
  1293
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1294
    (* existential closure over schematic variables *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1295
    (* (Term.indexname * Term.typ) list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1296
    val vars = sort_wrt (fst o fst) (map dest_Var (term_vars t))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1297
    (* Term.term *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1298
    val ex_closure = Library.foldl (fn (t', ((x, i), T)) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1299
      (HOLogic.exists_const T) $
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1300
        Abs (x, T, abstract_over (Var ((x, i), T), t')))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1301
      (t, vars)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1302
    (* Note: If 't' is of type 'propT' (rather than 'boolT'), applying   *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1303
    (* 'HOLogic.exists_const' is not type-correct.  However, this is not *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1304
    (* really a problem as long as 'find_model' still interprets the     *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1305
    (* resulting term correctly, without checking its type.              *)
21556
e0ffb2d13f9f outermost universal quantifiers are stripped
webertj
parents: 21267
diff changeset
  1306
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1307
    (* replace outermost universally quantified variables by Free's:     *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1308
    (* refuting a term with Free's is generally faster than refuting a   *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1309
    (* term with (nested) quantifiers, because quantifiers are expanded, *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1310
    (* while the SAT solver searches for an interpretation for Free's.   *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1311
    (* Also we get more information back that way, namely an             *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1312
    (* interpretation which includes values for the (formerly)           *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1313
    (* quantified variables.                                             *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1314
    (* maps  !!x1...xn. !xk...xm. t   to   t  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1315
    fun strip_all_body (Const ("all", _) $ Abs (_, _, t)) = strip_all_body t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1316
      | strip_all_body (Const ("Trueprop", _) $ t)        = strip_all_body t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1317
      | strip_all_body (Const ("All", _) $ Abs (_, _, t)) = strip_all_body t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1318
      | strip_all_body t                                  = t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1319
    (* maps  !!x1...xn. !xk...xm. t   to   [x1, ..., xn, xk, ..., xm]  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1320
    fun strip_all_vars (Const ("all", _) $ Abs (a, T, t)) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1321
      (a, T) :: strip_all_vars t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1322
      | strip_all_vars (Const ("Trueprop", _) $ t)        =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1323
      strip_all_vars t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1324
      | strip_all_vars (Const ("All", _) $ Abs (a, T, t)) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1325
      (a, T) :: strip_all_vars t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1326
      | strip_all_vars t                                  =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1327
      [] : (string * typ) list
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1328
    val strip_t = strip_all_body ex_closure
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1329
    val frees   = Term.rename_wrt_term strip_t (strip_all_vars ex_closure)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1330
    val subst_t = Term.subst_bounds (map Free frees, strip_t)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1331
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1332
    find_model thy (actual_params thy params) subst_t true
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1333
  end;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1334
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1335
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1336
(* refute_subgoal: calls 'refute_term' on a specific subgoal                 *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1337
(* params        : list of '(name, value)' pairs used to override default    *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1338
(*                 parameters                                                *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1339
(* subgoal       : 0-based index specifying the subgoal number               *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1340
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1341
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1342
  (* theory -> (string * string) list -> Thm.thm -> int -> unit *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1343
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1344
  fun refute_subgoal thy params thm subgoal =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1345
    refute_term thy params (List.nth (Thm.prems_of thm, subgoal));
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1346
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1347
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1348
(* ------------------------------------------------------------------------- *)
15292
09e218879265 minor changes (comments/code refactoring)
webertj
parents: 15283
diff changeset
  1349
(* INTERPRETERS: Auxiliary Functions                                         *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1350
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1351
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1352
(* ------------------------------------------------------------------------- *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1353
(* make_constants: returns all interpretations that have the same tree       *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1354
(*                 structure as 'intr', but consist of unit vectors with     *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1355
(*                 'True'/'False' only (no Boolean variables)                *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1356
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1357
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1358
  (* interpretation -> interpretation list *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1359
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1360
  fun make_constants intr =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1361
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1362
    (* returns a list with all unit vectors of length n *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1363
    (* int -> interpretation list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1364
    fun unit_vectors n =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1365
    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1366
      (* returns the k-th unit vector of length n *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1367
      (* int * int -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1368
      fun unit_vector (k,n) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1369
        Leaf ((replicate (k-1) False) @ (True :: (replicate (n-k) False)))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1370
      (* int -> interpretation list -> interpretation list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1371
      fun unit_vectors_acc k vs =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1372
        if k>n then [] else (unit_vector (k,n))::(unit_vectors_acc (k+1) vs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1373
    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1374
      unit_vectors_acc 1 []
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1375
    end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1376
    (* returns a list of lists, each one consisting of n (possibly *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1377
    (* identical) elements from 'xs'                               *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1378
    (* int -> 'a list -> 'a list list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1379
    fun pick_all 1 xs =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1380
      map single xs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1381
      | pick_all n xs =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1382
      let val rec_pick = pick_all (n-1) xs in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1383
        Library.foldl (fn (acc, x) => map (cons x) rec_pick @ acc) ([], xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1384
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1385
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1386
    case intr of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1387
      Leaf xs => unit_vectors (length xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1388
    | Node xs => map (fn xs' => Node xs') (pick_all (length xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1389
      (make_constants (hd xs)))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1390
  end;
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1391
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1392
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1393
(* size_of_type: returns the number of constants in a type (i.e. 'length     *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1394
(*               (make_constants intr)', but implemented more efficiently)   *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1395
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1396
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1397
  (* interpretation -> int *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1398
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1399
  fun size_of_type intr =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1400
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1401
    (* power (a, b) computes a^b, for a>=0, b>=0 *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1402
    (* int * int -> int *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1403
    fun power (a, 0) = 1
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1404
      | power (a, 1) = a
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1405
      | power (a, b) = let val ab = power(a, b div 2) in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1406
        ab * ab * power(a, b mod 2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1407
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1408
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1409
    case intr of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1410
      Leaf xs => length xs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1411
    | Node xs => power (size_of_type (hd xs), length xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1412
  end;
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1413
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1414
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1415
(* TT/FF: interpretations that denote "true" or "false", respectively        *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1416
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1417
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1418
  (* interpretation *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1419
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1420
  val TT = Leaf [True, False];
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1421
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1422
  val FF = Leaf [False, True];
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1423
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1424
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1425
(* make_equality: returns an interpretation that denotes (extensional)       *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1426
(*                equality of two interpretations                            *)
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1427
(* - two interpretations are 'equal' iff they are both defined and denote    *)
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1428
(*   the same value                                                          *)
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1429
(* - two interpretations are 'not_equal' iff they are both defined at least  *)
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1430
(*   partially, and a defined part denotes different values                  *)
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1431
(* - a completely undefined interpretation is neither 'equal' nor            *)
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1432
(*   'not_equal' to another interpretation                                   *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1433
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1434
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1435
  (* We could in principle represent '=' on a type T by a particular        *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1436
  (* interpretation.  However, the size of that interpretation is quadratic *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1437
  (* in the size of T.  Therefore comparing the interpretations 'i1' and    *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1438
  (* 'i2' directly is more efficient than constructing the interpretation   *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1439
  (* for equality on T first, and "applying" this interpretation to 'i1'    *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1440
  (* and 'i2' in the usual way (cf. 'interpretation_apply') then.           *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1441
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1442
  (* interpretation * interpretation -> interpretation *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1443
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1444
  fun make_equality (i1, i2) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1445
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1446
    (* interpretation * interpretation -> prop_formula *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1447
    fun equal (i1, i2) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1448
      (case i1 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1449
        Leaf xs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1450
        (case i2 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1451
          Leaf ys => PropLogic.dot_product (xs, ys)  (* defined and equal *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1452
        | Node _  => raise REFUTE ("make_equality",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1453
          "second interpretation is higher"))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1454
      | Node xs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1455
        (case i2 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1456
          Leaf _  => raise REFUTE ("make_equality",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1457
          "first interpretation is higher")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1458
        | Node ys => PropLogic.all (map equal (xs ~~ ys))))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1459
    (* interpretation * interpretation -> prop_formula *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1460
    fun not_equal (i1, i2) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1461
      (case i1 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1462
        Leaf xs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1463
        (case i2 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1464
          (* defined and not equal *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1465
          Leaf ys => PropLogic.all ((PropLogic.exists xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1466
          :: (PropLogic.exists ys)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1467
          :: (map (fn (x,y) => SOr (SNot x, SNot y)) (xs ~~ ys)))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1468
        | Node _  => raise REFUTE ("make_equality",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1469
          "second interpretation is higher"))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1470
      | Node xs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1471
        (case i2 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1472
          Leaf _  => raise REFUTE ("make_equality",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1473
          "first interpretation is higher")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1474
        | Node ys => PropLogic.exists (map not_equal (xs ~~ ys))))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1475
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1476
    (* a value may be undefined; therefore 'not_equal' is not just the *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1477
    (* negation of 'equal'                                             *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1478
    Leaf [equal (i1, i2), not_equal (i1, i2)]
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1479
  end;
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1480
15292
09e218879265 minor changes (comments/code refactoring)
webertj
parents: 15283
diff changeset
  1481
(* ------------------------------------------------------------------------- *)
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1482
(* make_def_equality: returns an interpretation that denotes (extensional)   *)
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1483
(*                    equality of two interpretations                        *)
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1484
(* This function treats undefined/partially defined interpretations          *)
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1485
(* different from 'make_equality': two undefined interpretations are         *)
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1486
(* considered equal, while a defined interpretation is considered not equal  *)
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1487
(* to an undefined interpretation.                                           *)
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1488
(* ------------------------------------------------------------------------- *)
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1489
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1490
  (* interpretation * interpretation -> interpretation *)
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1491
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1492
  fun make_def_equality (i1, i2) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1493
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1494
    (* interpretation * interpretation -> prop_formula *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1495
    fun equal (i1, i2) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1496
      (case i1 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1497
        Leaf xs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1498
        (case i2 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1499
          (* defined and equal, or both undefined *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1500
          Leaf ys => SOr (PropLogic.dot_product (xs, ys),
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1501
          SAnd (PropLogic.all (map SNot xs), PropLogic.all (map SNot ys)))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1502
        | Node _  => raise REFUTE ("make_def_equality",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1503
          "second interpretation is higher"))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1504
      | Node xs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1505
        (case i2 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1506
          Leaf _  => raise REFUTE ("make_def_equality",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1507
          "first interpretation is higher")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1508
        | Node ys => PropLogic.all (map equal (xs ~~ ys))))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1509
    (* interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1510
    val eq = equal (i1, i2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1511
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1512
    Leaf [eq, SNot eq]
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1513
  end;
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1514
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1515
(* ------------------------------------------------------------------------- *)
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1516
(* interpretation_apply: returns an interpretation that denotes the result   *)
22092
ab3dfcef6489 reformatted to 80 chars/line
webertj
parents: 22055
diff changeset
  1517
(*                       of applying the function denoted by 'i1' to the     *)
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1518
(*                       argument denoted by 'i2'                            *)
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1519
(* ------------------------------------------------------------------------- *)
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1520
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1521
  (* interpretation * interpretation -> interpretation *)
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1522
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1523
  fun interpretation_apply (i1, i2) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1524
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1525
    (* interpretation * interpretation -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1526
    fun interpretation_disjunction (tr1,tr2) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1527
      tree_map (fn (xs,ys) => map (fn (x,y) => SOr(x,y)) (xs ~~ ys))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1528
        (tree_pair (tr1,tr2))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1529
    (* prop_formula * interpretation -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1530
    fun prop_formula_times_interpretation (fm,tr) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1531
      tree_map (map (fn x => SAnd (fm,x))) tr
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1532
    (* prop_formula list * interpretation list -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1533
    fun prop_formula_list_dot_product_interpretation_list ([fm],[tr]) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1534
      prop_formula_times_interpretation (fm,tr)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1535
      | prop_formula_list_dot_product_interpretation_list (fm::fms,tr::trees) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1536
      interpretation_disjunction (prop_formula_times_interpretation (fm,tr),
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1537
        prop_formula_list_dot_product_interpretation_list (fms,trees))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1538
      | prop_formula_list_dot_product_interpretation_list (_,_) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1539
      raise REFUTE ("interpretation_apply", "empty list (in dot product)")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1540
    (* concatenates 'x' with every list in 'xss', returning a new list of *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1541
    (* lists                                                              *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1542
    (* 'a -> 'a list list -> 'a list list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1543
    fun cons_list x xss =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1544
      map (cons x) xss
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1545
    (* returns a list of lists, each one consisting of one element from each *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1546
    (* element of 'xss'                                                      *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1547
    (* 'a list list -> 'a list list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1548
    fun pick_all [xs] =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1549
      map single xs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1550
      | pick_all (xs::xss) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1551
      let val rec_pick = pick_all xss in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1552
        Library.foldl (fn (acc, x) => (cons_list x rec_pick) @ acc) ([], xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1553
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1554
      | pick_all _ =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1555
      raise REFUTE ("interpretation_apply", "empty list (in pick_all)")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1556
    (* interpretation -> prop_formula list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1557
    fun interpretation_to_prop_formula_list (Leaf xs) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1558
      xs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1559
      | interpretation_to_prop_formula_list (Node trees) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1560
      map PropLogic.all (pick_all
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1561
        (map interpretation_to_prop_formula_list trees))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1562
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1563
    case i1 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1564
      Leaf _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1565
      raise REFUTE ("interpretation_apply", "first interpretation is a leaf")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1566
    | Node xs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1567
      prop_formula_list_dot_product_interpretation_list
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1568
        (interpretation_to_prop_formula_list i2, xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1569
  end;
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1570
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1571
(* ------------------------------------------------------------------------- *)
15292
09e218879265 minor changes (comments/code refactoring)
webertj
parents: 15283
diff changeset
  1572
(* eta_expand: eta-expands a term 't' by adding 'i' lambda abstractions      *)
09e218879265 minor changes (comments/code refactoring)
webertj
parents: 15283
diff changeset
  1573
(* ------------------------------------------------------------------------- *)
09e218879265 minor changes (comments/code refactoring)
webertj
parents: 15283
diff changeset
  1574
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1575
  (* Term.term -> int -> Term.term *)
15292
09e218879265 minor changes (comments/code refactoring)
webertj
parents: 15283
diff changeset
  1576
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1577
  fun eta_expand t i =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1578
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1579
    val Ts = Term.binder_types (Term.fastype_of t)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1580
    val t' = Term.incr_boundvars i t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1581
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1582
    foldr (fn (T, term) => Abs ("<eta_expand>", T, term))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1583
      (Term.list_comb (t', map Bound (i-1 downto 0))) (List.take (Ts, i))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1584
  end;
15292
09e218879265 minor changes (comments/code refactoring)
webertj
parents: 15283
diff changeset
  1585
15335
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
  1586
(* ------------------------------------------------------------------------- *)
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
  1587
(* sum: returns the sum of a list 'xs' of integers                           *)
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
  1588
(* ------------------------------------------------------------------------- *)
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
  1589
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1590
  (* int list -> int *)
15335
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
  1591
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1592
  fun sum xs = foldl op+ 0 xs;
15335
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
  1593
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
  1594
(* ------------------------------------------------------------------------- *)
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
  1595
(* product: returns the product of a list 'xs' of integers                   *)
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
  1596
(* ------------------------------------------------------------------------- *)
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
  1597
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1598
  (* int list -> int *)
15335
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
  1599
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1600
  fun product xs = foldl op* 1 xs;
15335
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
  1601
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
  1602
(* ------------------------------------------------------------------------- *)
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1603
(* size_of_dtyp: the size of (an initial fragment of) an inductive data type *)
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1604
(*               is the sum (over its constructors) of the product (over     *)
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1605
(*               their arguments) of the size of the argument types          *)
15335
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
  1606
(* ------------------------------------------------------------------------- *)
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
  1607
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1608
  (* theory -> (Term.typ * int) list -> DatatypeAux.descr ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1609
    (DatatypeAux.dtyp * Term.typ) list ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1610
    (string * DatatypeAux.dtyp list) list -> int *)
15335
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
  1611
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1612
  fun size_of_dtyp thy typ_sizes descr typ_assoc constructors =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1613
    sum (map (fn (_, dtyps) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1614
      product (map (fn dtyp =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1615
        let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1616
          val T         = typ_of_dtyp descr typ_assoc dtyp
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1617
          val (i, _, _) = interpret thy (typ_sizes, [])
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1618
            {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1619
            (Free ("dummy", T))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1620
        in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1621
          size_of_type i
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1622
        end) dtyps)) constructors);
15335
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
  1623
15292
09e218879265 minor changes (comments/code refactoring)
webertj
parents: 15283
diff changeset
  1624
09e218879265 minor changes (comments/code refactoring)
webertj
parents: 15283
diff changeset
  1625
(* ------------------------------------------------------------------------- *)
09e218879265 minor changes (comments/code refactoring)
webertj
parents: 15283
diff changeset
  1626
(* INTERPRETERS: Actual Interpreters                                         *)
09e218879265 minor changes (comments/code refactoring)
webertj
parents: 15283
diff changeset
  1627
(* ------------------------------------------------------------------------- *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1628
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1629
  (* theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1630
    (interpretation * model * arguments) option *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1631
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1632
  (* simply typed lambda calculus: Isabelle's basic term syntax, with type *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1633
  (* variables, function types, and propT                                  *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1634
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1635
  fun stlc_interpreter thy model args t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1636
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1637
    val (typs, terms)                                   = model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1638
    val {maxvars, def_eq, next_idx, bounds, wellformed} = args
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1639
    (* Term.typ -> (interpretation * model * arguments) option *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1640
    fun interpret_groundterm T =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1641
    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1642
      (* unit -> (interpretation * model * arguments) option *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1643
      fun interpret_groundtype () =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1644
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1645
        (* the model must specify a size for ground types *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1646
        val size = (if T = Term.propT then 2 else lookup typs T)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1647
        val next = next_idx+size
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1648
        (* check if 'maxvars' is large enough *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1649
        val _    = (if next-1>maxvars andalso maxvars>0 then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1650
          raise MAXVARS_EXCEEDED else ())
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1651
        (* prop_formula list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1652
        val fms  = map BoolVar (next_idx upto (next_idx+size-1))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1653
        (* interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1654
        val intr = Leaf fms
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1655
        (* prop_formula list -> prop_formula *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1656
        fun one_of_two_false []      = True
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1657
          | one_of_two_false (x::xs) = SAnd (PropLogic.all (map (fn x' =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1658
          SOr (SNot x, SNot x')) xs), one_of_two_false xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1659
        (* prop_formula *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1660
        val wf   = one_of_two_false fms
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1661
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1662
        (* extend the model, increase 'next_idx', add well-formedness *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1663
        (* condition                                                  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1664
        SOME (intr, (typs, (t, intr)::terms), {maxvars = maxvars,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1665
          def_eq = def_eq, next_idx = next, bounds = bounds,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1666
          wellformed = SAnd (wellformed, wf)})
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1667
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1668
    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1669
      case T of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1670
        Type ("fun", [T1, T2]) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1671
        let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1672
          (* we create 'size_of_type (interpret (... T1))' different copies *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1673
          (* of the interpretation for 'T2', which are then combined into a *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1674
          (* single new interpretation                                      *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1675
          val (i1, _, _) = interpret thy model {maxvars=0, def_eq=false,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1676
            next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T1))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1677
          (* make fresh copies, with different variable indices *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1678
          (* 'idx': next variable index                         *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1679
          (* 'n'  : number of copies                            *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1680
          (* int -> int -> (int * interpretation list * prop_formula *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1681
          fun make_copies idx 0 =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1682
            (idx, [], True)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1683
            | make_copies idx n =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1684
            let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1685
              val (copy, _, new_args) = interpret thy (typs, [])
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1686
                {maxvars = maxvars, def_eq = false, next_idx = idx,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1687
                bounds = [], wellformed = True} (Free ("dummy", T2))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1688
              val (idx', copies, wf') = make_copies (#next_idx new_args) (n-1)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1689
            in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1690
              (idx', copy :: copies, SAnd (#wellformed new_args, wf'))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1691
            end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1692
          val (next, copies, wf) = make_copies next_idx (size_of_type i1)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1693
          (* combine copies into a single interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1694
          val intr = Node copies
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1695
        in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1696
          (* extend the model, increase 'next_idx', add well-formedness *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1697
          (* condition                                                  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1698
          SOME (intr, (typs, (t, intr)::terms), {maxvars = maxvars,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1699
            def_eq = def_eq, next_idx = next, bounds = bounds,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1700
            wellformed = SAnd (wellformed, wf)})
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1701
        end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1702
      | Type _  => interpret_groundtype ()
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1703
      | TFree _ => interpret_groundtype ()
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1704
      | TVar  _ => interpret_groundtype ()
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1705
    end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1706
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1707
    case AList.lookup (op =) terms t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1708
      SOME intr =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1709
      (* return an existing interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1710
      SOME (intr, model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1711
    | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1712
      (case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1713
        Const (_, T)     =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1714
        interpret_groundterm T
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1715
      | Free (_, T)      =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1716
        interpret_groundterm T
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1717
      | Var (_, T)       =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1718
        interpret_groundterm T
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1719
      | Bound i          =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1720
        SOME (List.nth (#bounds args, i), model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1721
      | Abs (x, T, body) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1722
        let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1723
          (* create all constants of type 'T' *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1724
          val (i, _, _) = interpret thy model {maxvars=0, def_eq=false,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1725
            next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1726
          val constants = make_constants i
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1727
          (* interpret the 'body' separately for each constant *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1728
          val ((model', args'), bodies) = foldl_map
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1729
            (fn ((m, a), c) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1730
              let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1731
                (* add 'c' to 'bounds' *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1732
                val (i', m', a') = interpret thy m {maxvars = #maxvars a,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1733
                  def_eq = #def_eq a, next_idx = #next_idx a,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1734
                  bounds = (c :: #bounds a), wellformed = #wellformed a} body
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1735
              in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1736
                (* keep the new model m' and 'next_idx' and 'wellformed', *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1737
                (* but use old 'bounds'                                   *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1738
                ((m', {maxvars = maxvars, def_eq = def_eq,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1739
                  next_idx = #next_idx a', bounds = bounds,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1740
                  wellformed = #wellformed a'}), i')
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1741
              end)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1742
            ((model, args), constants)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1743
        in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1744
          SOME (Node bodies, model', args')
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1745
        end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1746
      | t1 $ t2          =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1747
        let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1748
          (* interpret 't1' and 't2' separately *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1749
          val (intr1, model1, args1) = interpret thy model args t1
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1750
          val (intr2, model2, args2) = interpret thy model1 args1 t2
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1751
        in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1752
          SOME (interpretation_apply (intr1, intr2), model2, args2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1753
        end)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1754
  end;
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1755
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1756
  (* theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1757
    (interpretation * model * arguments) option *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1758
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1759
  fun Pure_interpreter thy model args t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1760
    case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1761
      Const ("all", _) $ t1 =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1762
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1763
        val (i, m, a) = interpret thy model args t1
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1764
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1765
        case i of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1766
          Node xs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1767
          (* 3-valued logic *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1768
          let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1769
            val fmTrue  = PropLogic.all (map toTrue xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1770
            val fmFalse = PropLogic.exists (map toFalse xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1771
          in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1772
            SOME (Leaf [fmTrue, fmFalse], m, a)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1773
          end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1774
        | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1775
          raise REFUTE ("Pure_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1776
            "\"all\" is followed by a non-function")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1777
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1778
    | Const ("all", _) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1779
      SOME (interpret thy model args (eta_expand t 1))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1780
    | Const ("==", _) $ t1 $ t2 =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1781
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1782
        val (i1, m1, a1) = interpret thy model args t1
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1783
        val (i2, m2, a2) = interpret thy m1 a1 t2
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1784
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1785
        (* we use either 'make_def_equality' or 'make_equality' *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1786
        SOME ((if #def_eq args then make_def_equality else make_equality)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1787
          (i1, i2), m2, a2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1788
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1789
    | Const ("==", _) $ t1 =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1790
      SOME (interpret thy model args (eta_expand t 1))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1791
    | Const ("==", _) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1792
      SOME (interpret thy model args (eta_expand t 2))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1793
    | Const ("==>", _) $ t1 $ t2 =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1794
      (* 3-valued logic *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1795
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1796
        val (i1, m1, a1) = interpret thy model args t1
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1797
        val (i2, m2, a2) = interpret thy m1 a1 t2
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1798
        val fmTrue       = PropLogic.SOr (toFalse i1, toTrue i2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1799
        val fmFalse      = PropLogic.SAnd (toTrue i1, toFalse i2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1800
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1801
        SOME (Leaf [fmTrue, fmFalse], m2, a2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1802
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1803
    | Const ("==>", _) $ t1 =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1804
      SOME (interpret thy model args (eta_expand t 1))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1805
    | Const ("==>", _) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1806
      SOME (interpret thy model args (eta_expand t 2))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1807
    | _ => NONE;
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1808
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1809
  (* theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1810
    (interpretation * model * arguments) option *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1811
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1812
  fun HOLogic_interpreter thy model args t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1813
  (* Providing interpretations directly is more efficient than unfolding the *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1814
  (* logical constants.  In HOL however, logical constants can themselves be *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1815
  (* arguments.  They are then translated using eta-expansion.               *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1816
    case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1817
      Const ("Trueprop", _) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1818
      SOME (Node [TT, FF], model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1819
    | Const ("Not", _) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1820
      SOME (Node [FF, TT], model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1821
    (* redundant, since 'True' is also an IDT constructor *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1822
    | Const ("True", _) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1823
      SOME (TT, model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1824
    (* redundant, since 'False' is also an IDT constructor *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1825
    | Const ("False", _) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1826
      SOME (FF, model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1827
    | Const ("All", _) $ t1 =>  (* similar to "all" (Pure) *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1828
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1829
        val (i, m, a) = interpret thy model args t1
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1830
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1831
        case i of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1832
          Node xs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1833
          (* 3-valued logic *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1834
          let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1835
            val fmTrue  = PropLogic.all (map toTrue xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1836
            val fmFalse = PropLogic.exists (map toFalse xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1837
          in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1838
            SOME (Leaf [fmTrue, fmFalse], m, a)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1839
          end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1840
        | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1841
          raise REFUTE ("HOLogic_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1842
            "\"All\" is followed by a non-function")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1843
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1844
    | Const ("All", _) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1845
      SOME (interpret thy model args (eta_expand t 1))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1846
    | Const ("Ex", _) $ t1 =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1847
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1848
        val (i, m, a) = interpret thy model args t1
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1849
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1850
        case i of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1851
          Node xs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1852
          (* 3-valued logic *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1853
          let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1854
            val fmTrue  = PropLogic.exists (map toTrue xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1855
            val fmFalse = PropLogic.all (map toFalse xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1856
          in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1857
            SOME (Leaf [fmTrue, fmFalse], m, a)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1858
          end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1859
        | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1860
          raise REFUTE ("HOLogic_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1861
            "\"Ex\" is followed by a non-function")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1862
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1863
    | Const ("Ex", _) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1864
      SOME (interpret thy model args (eta_expand t 1))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1865
    | Const ("op =", _) $ t1 $ t2 =>  (* similar to "==" (Pure) *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1866
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1867
        val (i1, m1, a1) = interpret thy model args t1
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1868
        val (i2, m2, a2) = interpret thy m1 a1 t2
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1869
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1870
        SOME (make_equality (i1, i2), m2, a2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1871
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1872
    | Const ("op =", _) $ t1 =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1873
      SOME (interpret thy model args (eta_expand t 1))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1874
    | Const ("op =", _) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1875
      SOME (interpret thy model args (eta_expand t 2))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1876
    | Const ("op &", _) $ t1 $ t2 =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1877
      (* 3-valued logic *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1878
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1879
        val (i1, m1, a1) = interpret thy model args t1
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1880
        val (i2, m2, a2) = interpret thy m1 a1 t2
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1881
        val fmTrue       = PropLogic.SAnd (toTrue i1, toTrue i2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1882
        val fmFalse      = PropLogic.SOr (toFalse i1, toFalse i2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1883
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1884
        SOME (Leaf [fmTrue, fmFalse], m2, a2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1885
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1886
    | Const ("op &", _) $ t1 =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1887
      SOME (interpret thy model args (eta_expand t 1))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1888
    | Const ("op &", _) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1889
      SOME (interpret thy model args (eta_expand t 2))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1890
      (* this would make "undef" propagate, even for formulae like *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1891
      (* "False & undef":                                          *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1892
      (* SOME (Node [Node [TT, FF], Node [FF, FF]], model, args) *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1893
    | Const ("op |", _) $ t1 $ t2 =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1894
      (* 3-valued logic *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1895
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1896
        val (i1, m1, a1) = interpret thy model args t1
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1897
        val (i2, m2, a2) = interpret thy m1 a1 t2
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1898
        val fmTrue       = PropLogic.SOr (toTrue i1, toTrue i2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1899
        val fmFalse      = PropLogic.SAnd (toFalse i1, toFalse i2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1900
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1901
        SOME (Leaf [fmTrue, fmFalse], m2, a2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1902
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1903
    | Const ("op |", _) $ t1 =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1904
      SOME (interpret thy model args (eta_expand t 1))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1905
    | Const ("op |", _) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1906
      SOME (interpret thy model args (eta_expand t 2))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1907
      (* this would make "undef" propagate, even for formulae like *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1908
      (* "True | undef":                                           *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1909
      (* SOME (Node [Node [TT, TT], Node [TT, FF]], model, args) *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1910
    | Const ("op -->", _) $ t1 $ t2 =>  (* similar to "==>" (Pure) *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1911
      (* 3-valued logic *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1912
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1913
        val (i1, m1, a1) = interpret thy model args t1
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1914
        val (i2, m2, a2) = interpret thy m1 a1 t2
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1915
        val fmTrue       = PropLogic.SOr (toFalse i1, toTrue i2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1916
        val fmFalse      = PropLogic.SAnd (toTrue i1, toFalse i2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1917
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1918
        SOME (Leaf [fmTrue, fmFalse], m2, a2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1919
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1920
    | Const ("op -->", _) $ t1 =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1921
      SOME (interpret thy model args (eta_expand t 1))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1922
    | Const ("op -->", _) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1923
      SOME (interpret thy model args (eta_expand t 2))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1924
      (* this would make "undef" propagate, even for formulae like *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1925
      (* "False --> undef":                                        *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1926
      (* SOME (Node [Node [TT, FF], Node [TT, TT]], model, args) *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1927
    | _ => NONE;
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1928
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1929
  (* theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1930
    (interpretation * model * arguments) option *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1931
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1932
  fun set_interpreter thy model args t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1933
  (* "T set" is isomorphic to "T --> bool" *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1934
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1935
    val (typs, terms) = model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1936
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1937
    case AList.lookup (op =) terms t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1938
      SOME intr =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1939
      (* return an existing interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1940
      SOME (intr, model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1941
    | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1942
      (case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1943
        Free (x, Type ("set", [T])) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1944
        let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1945
          val (intr, _, args') =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1946
            interpret thy (typs, []) args (Free (x, T --> HOLogic.boolT))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1947
        in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1948
          SOME (intr, (typs, (t, intr)::terms), args')
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1949
        end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1950
      | Var ((x, i), Type ("set", [T])) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1951
        let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1952
          val (intr, _, args') =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1953
            interpret thy (typs, []) args (Var ((x,i), T --> HOLogic.boolT))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1954
        in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1955
          SOME (intr, (typs, (t, intr)::terms), args')
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1956
        end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1957
      | Const (s, Type ("set", [T])) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1958
        let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1959
          val (intr, _, args') =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1960
            interpret thy (typs, []) args (Const (s, T --> HOLogic.boolT))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1961
        in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1962
          SOME (intr, (typs, (t, intr)::terms), args')
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1963
        end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1964
      (* 'Collect' == identity *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1965
      | Const ("Collect", _) $ t1 =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1966
        SOME (interpret thy model args t1)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1967
      | Const ("Collect", _) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1968
        SOME (interpret thy model args (eta_expand t 1))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1969
      (* 'op :' == application *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1970
      | Const ("op :", _) $ t1 $ t2 =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1971
        SOME (interpret thy model args (t2 $ t1))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1972
      | Const ("op :", _) $ t1 =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1973
        SOME (interpret thy model args (eta_expand t 1))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1974
      | Const ("op :", _) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1975
        SOME (interpret thy model args (eta_expand t 2))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1976
      | _ => NONE)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1977
  end;
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1978
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1979
  (* theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1980
    (interpretation * model * arguments) option *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  1981
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1982
  (* interprets variables and constants whose type is an IDT; *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1983
  (* constructors of IDTs however are properly interpreted by *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1984
  (* 'IDT_constructor_interpreter'                            *)
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  1985
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1986
  fun IDT_interpreter thy model args t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1987
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1988
    val (typs, terms) = model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1989
    (* Term.typ -> (interpretation * model * arguments) option *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1990
    fun interpret_term (Type (s, Ts)) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1991
      (case DatatypePackage.get_datatype thy s of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1992
        SOME info =>  (* inductive datatype *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1993
        let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1994
          (* int option -- only recursive IDTs have an associated depth *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1995
          val depth = AList.lookup (op =) typs (Type (s, Ts))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1996
        in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1997
          (* termination condition to avoid infinite recursion *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1998
          if depth = (SOME 0) then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1999
            (* return a leaf of size 0 *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2000
            SOME (Leaf [], model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2001
          else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2002
            let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2003
              val index               = #index info
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2004
              val descr               = #descr info
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2005
              val (_, dtyps, constrs) = lookup descr index
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2006
              val typ_assoc           = dtyps ~~ Ts
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2007
              (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2008
              val _ = (if Library.exists (fn d =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2009
                  case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2010
                then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2011
                  raise REFUTE ("IDT_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2012
                    "datatype argument (for type "
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2013
                    ^ Sign.string_of_typ thy (Type (s, Ts))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2014
                    ^ ") is not a variable")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2015
                else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2016
                  ())
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2017
              (* if the model specifies a depth for the current type, *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2018
              (* decrement it to avoid infinite recursion             *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2019
              val typs'    = case depth of NONE => typs | SOME n =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2020
                AList.update (op =) (Type (s, Ts), n-1) typs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2021
              (* recursively compute the size of the datatype *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2022
              val size     = size_of_dtyp thy typs' descr typ_assoc constrs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2023
              val next_idx = #next_idx args
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2024
              val next     = next_idx+size
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2025
              (* check if 'maxvars' is large enough *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2026
              val _        = (if next-1 > #maxvars args andalso
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2027
                #maxvars args > 0 then raise MAXVARS_EXCEEDED else ())
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2028
              (* prop_formula list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2029
              val fms      = map BoolVar (next_idx upto (next_idx+size-1))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2030
              (* interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2031
              val intr     = Leaf fms
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2032
              (* prop_formula list -> prop_formula *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2033
              fun one_of_two_false []      = True
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2034
                | one_of_two_false (x::xs) = SAnd (PropLogic.all (map (fn x' =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2035
                SOr (SNot x, SNot x')) xs), one_of_two_false xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2036
              (* prop_formula *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2037
              val wf       = one_of_two_false fms
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2038
            in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2039
              (* extend the model, increase 'next_idx', add well-formedness *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2040
              (* condition                                                  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2041
              SOME (intr, (typs, (t, intr)::terms), {maxvars = #maxvars args,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2042
                def_eq = #def_eq args, next_idx = next, bounds = #bounds args,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2043
                wellformed = SAnd (#wellformed args, wf)})
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2044
            end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2045
        end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2046
      | NONE =>  (* not an inductive datatype *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2047
        NONE)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2048
      | interpret_term _ =  (* a (free or schematic) type variable *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2049
      NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2050
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2051
    case AList.lookup (op =) terms t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2052
      SOME intr =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2053
      (* return an existing interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2054
      SOME (intr, model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2055
    | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2056
      (case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2057
        Free (_, T)  => interpret_term T
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2058
      | Var (_, T)   => interpret_term T
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2059
      | Const (_, T) => interpret_term T
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2060
      | _            => NONE)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2061
  end;
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  2062
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2063
  (* theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2064
    (interpretation * model * arguments) option *)
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  2065
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2066
  fun IDT_constructor_interpreter thy model args t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2067
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2068
    val (typs, terms) = model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2069
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2070
    case AList.lookup (op =) terms t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2071
      SOME intr =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2072
      (* return an existing interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2073
      SOME (intr, model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2074
    | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2075
      (case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2076
        Const (s, T) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2077
        (case body_type T of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2078
          Type (s', Ts') =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2079
          (case DatatypePackage.get_datatype thy s' of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2080
            SOME info =>  (* body type is an inductive datatype *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2081
            let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2082
              val index               = #index info
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2083
              val descr               = #descr info
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2084
              val (_, dtyps, constrs) = lookup descr index
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2085
              val typ_assoc           = dtyps ~~ Ts'
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2086
              (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2087
              val _ = (if Library.exists (fn d =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2088
                  case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2089
                then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2090
                  raise REFUTE ("IDT_constructor_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2091
                    "datatype argument (for type "
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2092
                    ^ Sign.string_of_typ thy (Type (s', Ts'))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2093
                    ^ ") is not a variable")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2094
                else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2095
                  ())
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2096
              (* split the constructors into those occuring before/after *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2097
              (* 'Const (s, T)'                                          *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2098
              val (constrs1, constrs2) = take_prefix (fn (cname, ctypes) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2099
                not (cname = s andalso Sign.typ_instance thy (T,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2100
                  map (typ_of_dtyp descr typ_assoc) ctypes
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2101
                    ---> Type (s', Ts')))) constrs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2102
            in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2103
              case constrs2 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2104
                [] =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2105
                (* 'Const (s, T)' is not a constructor of this datatype *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2106
                NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2107
              | (_, ctypes)::cs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2108
                let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2109
                  (* compute the total size of the datatype (with the *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2110
                  (* current depth)                                   *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2111
                  val (i, _, _) = interpret thy (typs, []) {maxvars=0,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2112
                    def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2113
                    (Free ("dummy", Type (s', Ts')))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2114
                  val total     = size_of_type i
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2115
                  (* int option -- only /recursive/ IDTs have an associated *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2116
                  (*               depth                                    *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2117
                  val depth = AList.lookup (op =) typs (Type (s', Ts'))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2118
                  val typs' = (case depth of NONE => typs | SOME n =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2119
                    AList.update (op =) (Type (s', Ts'), n-1) typs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2120
                  (* returns an interpretation where everything is mapped to *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2121
                  (* "undefined"                                             *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2122
                  (* DatatypeAux.dtyp list -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2123
                  fun make_undef [] =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2124
                    Leaf (replicate total False)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2125
                    | make_undef (d::ds) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2126
                    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2127
                      (* compute the current size of the type 'd' *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2128
                      val T           = typ_of_dtyp descr typ_assoc d
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2129
                      val (i, _, _)   = interpret thy (typs, []) {maxvars=0,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2130
                        def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2131
                        (Free ("dummy", T))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2132
                      val size        = size_of_type i
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2133
                    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2134
                      Node (replicate size (make_undef ds))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2135
                    end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2136
                  (* returns the interpretation for a constructor at depth 1 *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2137
                  (* int * DatatypeAux.dtyp list -> int * interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2138
                  fun make_constr (offset, []) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2139
                    if offset<total then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2140
                      (offset+1, Leaf ((replicate offset False) @ True ::
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2141
                        (replicate (total-offset-1) False)))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2142
                    else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2143
                      raise REFUTE ("IDT_constructor_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2144
                        "offset >= total")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2145
                    | make_constr (offset, d::ds) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2146
                    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2147
                      (* compute the current and the old size of the type 'd' *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2148
                      val T           = typ_of_dtyp descr typ_assoc d
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2149
                      val (i, _, _)   = interpret thy (typs, []) {maxvars=0,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2150
                        def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2151
                        (Free ("dummy", T))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2152
                      val size        = size_of_type i
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2153
                      val (i', _, _)  = interpret thy (typs', []) {maxvars=0,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2154
                        def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2155
                        (Free ("dummy", T))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2156
                      val size'       = size_of_type i'
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2157
                      (* sanity check *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2158
                      val _           = if size < size' then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2159
                          raise REFUTE ("IDT_constructor_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2160
                            "current size is less than old size")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2161
                        else ()
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2162
                      (* int * interpretation list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2163
                      val (new_offset, intrs) = foldl_map make_constr
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2164
                        (offset, replicate size' ds)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2165
                      (* interpretation list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2166
                      val undefs = replicate (size - size') (make_undef ds)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2167
                    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2168
                      (* elements that exist at the previous depth are      *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2169
                      (* mapped to a defined value, while new elements are  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2170
                      (* mapped to "undefined" by the recursive constructor *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2171
                      (new_offset, Node (intrs @ undefs))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2172
                    end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2173
                  (* extends the interpretation for a constructor (both      *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2174
                  (* recursive and non-recursive) obtained at depth n (n>=1) *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2175
                  (* to depth n+1                                            *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2176
                  (* int * DatatypeAux.dtyp list * interpretation
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2177
                    -> int * interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2178
                  fun extend_constr (offset, [], Leaf xs) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2179
                    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2180
                      (* returns the k-th unit vector of length n *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2181
                      (* int * int -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2182
                      fun unit_vector (k, n) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2183
                        Leaf ((replicate (k-1) False) @ True ::
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2184
                          (replicate (n-k) False))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2185
                      (* int *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2186
                      val k = find_index_eq True xs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2187
                    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2188
                      if k=(~1) then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2189
                        (* if the element was mapped to "undefined" before, *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2190
                        (* map it to the value given by 'offset' now (and   *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2191
                        (* extend the length of the leaf)                   *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2192
                        (offset+1, unit_vector (offset+1, total))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2193
                      else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2194
                        (* if the element was already mapped to a defined  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2195
                        (* value, map it to the same value again, just     *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2196
                        (* extend the length of the leaf, do not increment *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2197
                        (* the 'offset'                                    *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2198
                        (offset, unit_vector (k+1, total))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2199
                    end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2200
                    | extend_constr (_, [], Node _) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2201
                    raise REFUTE ("IDT_constructor_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2202
                      "interpretation for constructor (with no arguments left)"
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2203
                      ^ " is a node")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2204
                    | extend_constr (offset, d::ds, Node xs) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2205
                    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2206
                      (* compute the size of the type 'd' *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2207
                      val T          = typ_of_dtyp descr typ_assoc d
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2208
                      val (i, _, _)  = interpret thy (typs, []) {maxvars=0,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2209
                        def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2210
                        (Free ("dummy", T))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2211
                      val size       = size_of_type i
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2212
                      (* sanity check *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2213
                      val _          = if size < length xs then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2214
                          raise REFUTE ("IDT_constructor_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2215
                            "new size of type is less than old size")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2216
                        else ()
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2217
                      (* extend the existing interpretations *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2218
                      (* int * interpretation list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2219
                      val (new_offset, intrs) = foldl_map (fn (off, i) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2220
                        extend_constr (off, ds, i)) (offset, xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2221
                      (* new elements of the type 'd' are mapped to *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2222
                      (* "undefined"                                *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2223
                      val undefs = replicate (size - length xs) (make_undef ds)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2224
                    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2225
                      (new_offset, Node (intrs @ undefs))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2226
                    end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2227
                    | extend_constr (_, d::ds, Leaf _) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2228
                    raise REFUTE ("IDT_constructor_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2229
                      "interpretation for constructor (with arguments left)"
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2230
                      ^ " is a leaf")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2231
                  (* returns 'true' iff the constructor has a recursive *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2232
                  (* argument                                           *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2233
                  (* DatatypeAux.dtyp list -> bool *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2234
                  fun is_rec_constr ds =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2235
                    Library.exists DatatypeAux.is_rec_type ds
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2236
                  (* constructors before 'Const (s, T)' generate elements of *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2237
                  (* the datatype                                            *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2238
                  val offset = size_of_dtyp thy typs' descr typ_assoc constrs1
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2239
                in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2240
                  case depth of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2241
                    NONE =>  (* equivalent to a depth of 1 *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2242
                    SOME (snd (make_constr (offset, ctypes)), model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2243
                  | SOME 0 =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2244
                    raise REFUTE ("IDT_constructor_interpreter", "depth is 0")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2245
                  | SOME 1 =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2246
                    SOME (snd (make_constr (offset, ctypes)), model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2247
                  | SOME n =>  (* n > 1 *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2248
                    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2249
                      (* interpret the constructor at depth-1 *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2250
                      val (iC, _, _) = interpret thy (typs', []) {maxvars=0,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2251
                        def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2252
                        (Const (s, T))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2253
                      (* elements generated by the constructor at depth-1 *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2254
                      (* must be added to 'offset'                        *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2255
                      (* interpretation -> int *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2256
                      fun number_of_defined_elements (Leaf xs) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2257
                        if find_index_eq True xs = (~1) then 0 else 1
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2258
                        | number_of_defined_elements (Node xs) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2259
                        sum (map number_of_defined_elements xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2260
                      (* int *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2261
                      val offset' = offset + number_of_defined_elements iC
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2262
                    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2263
                      SOME (snd (extend_constr (offset', ctypes, iC)), model,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2264
                        args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2265
                    end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2266
                end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2267
            end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2268
          | NONE =>  (* body type is not an inductive datatype *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2269
            NONE)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2270
        | _ =>  (* body type is a (free or schematic) type variable *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2271
          NONE)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2272
      | _ =>  (* term is not a constant *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2273
        NONE)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2274
  end;
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  2275
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2276
  (* theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2277
    (interpretation * model * arguments) option *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  2278
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2279
  (* Difficult code ahead.  Make sure you understand the                *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2280
  (* 'IDT_constructor_interpreter' and the order in which it enumerates *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2281
  (* elements of an IDT before you try to understand this function.     *)
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  2282
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2283
  fun IDT_recursion_interpreter thy model args t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2284
    (* careful: here we descend arbitrarily deep into 't', possibly before *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2285
    (* any other interpreter for atomic terms has had a chance to look at  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2286
    (* 't'                                                                 *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2287
    case strip_comb t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2288
      (Const (s, T), params) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2289
      (* iterate over all datatypes in 'thy' *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2290
      Symtab.fold (fn (_, info) => fn result =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2291
        case result of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2292
          SOME _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2293
          result  (* just keep 'result' *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2294
        | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2295
          if member (op =) (#rec_names info) s then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2296
            (* we do have a recursion operator of the datatype given by *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2297
            (* 'info', or of a mutually recursive datatype              *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2298
            let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2299
              val index              = #index info
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2300
              val descr              = #descr info
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2301
              val (dtname, dtyps, _) = lookup descr index
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2302
              (* number of all constructors, including those of different  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2303
              (* (mutually recursive) datatypes within the same descriptor *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2304
              (* 'descr'                                                   *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2305
              val mconstrs_count = sum (map (fn (_, (_, _, cs)) => length cs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2306
                descr)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2307
              val params_count   = length params
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2308
              (* the type of a recursion operator: *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2309
              (* [T1, ..., Tn, IDT] ---> Tresult   *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2310
              val IDT = List.nth (binder_types T, mconstrs_count)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2311
            in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2312
              if (fst o dest_Type) IDT <> dtname then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2313
                (* recursion operator of a mutually recursive datatype *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2314
                NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2315
              else if mconstrs_count < params_count then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2316
                (* too many actual parameters; for now we'll use the *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2317
                (* 'stlc_interpreter' to strip off one application   *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2318
                NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2319
              else if mconstrs_count > params_count then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2320
                (* too few actual parameters; we use eta expansion          *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2321
                (* Note that the resulting expansion of lambda abstractions *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2322
                (* by the 'stlc_interpreter' may be rather slow (depending  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2323
                (* on the argument types and the size of the IDT, of        *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2324
                (* course).                                                 *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2325
                SOME (interpret thy model args (eta_expand t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2326
                  (mconstrs_count - params_count)))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2327
              else  (* mconstrs_count = params_count *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2328
                let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2329
                  (* interpret each parameter separately *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2330
                  val ((model', args'), p_intrs) = foldl_map (fn ((m, a), p) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2331
                    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2332
                      val (i, m', a') = interpret thy m a p
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2333
                    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2334
                      ((m', a'), i)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2335
                    end) ((model, args), params)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2336
                  val (typs, _) = model'
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2337
                  val typ_assoc = dtyps ~~ (snd o dest_Type) IDT
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2338
                  (* interpret each constructor in the descriptor (including *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2339
                  (* those of mutually recursive datatypes)                  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2340
                  (* (int * interpretation list) list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2341
                  val mc_intrs = map (fn (idx, (_, _, cs)) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2342
                    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2343
                      val c_return_typ = typ_of_dtyp descr typ_assoc
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2344
                        (DatatypeAux.DtRec idx)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2345
                    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2346
                      (idx, map (fn (cname, cargs) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2347
                        (#1 o interpret thy (typs, []) {maxvars=0,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2348
                          def_eq=false, next_idx=1, bounds=[],
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2349
                          wellformed=True}) (Const (cname, map (typ_of_dtyp
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2350
                          descr typ_assoc) cargs ---> c_return_typ))) cs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2351
                    end) descr
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2352
                  (* the recursion operator is a function that maps every   *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2353
                  (* element of the inductive datatype (and of mutually     *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2354
                  (* recursive types) to an element of some result type; an *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2355
                  (* array entry of NONE means that the actual result has   *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2356
                  (* not been computed yet                                  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2357
                  (* (int * interpretation option Array.array) list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2358
                  val INTRS = map (fn (idx, _) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2359
                    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2360
                      val T         = typ_of_dtyp descr typ_assoc
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2361
                        (DatatypeAux.DtRec idx)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2362
                      val (i, _, _) = interpret thy (typs, []) {maxvars=0,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2363
                        def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2364
                        (Free ("dummy", T))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2365
                      val size      = size_of_type i
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2366
                    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2367
                      (idx, Array.array (size, NONE))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2368
                    end) descr
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2369
                  (* takes an interpretation, and if some leaf of this     *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2370
                  (* interpretation is the 'elem'-th element of the type,  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2371
                  (* the indices of the arguments leading to this leaf are *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2372
                  (* returned                                              *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2373
                  (* interpretation -> int -> int list option *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2374
                  fun get_args (Leaf xs) elem =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2375
                    if find_index_eq True xs = elem then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2376
                      SOME []
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2377
                    else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2378
                      NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2379
                    | get_args (Node xs) elem =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2380
                    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2381
                      (* interpretation * int -> int list option *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2382
                      fun search ([], _) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2383
                        NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2384
                        | search (x::xs, n) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2385
                        (case get_args x elem of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2386
                          SOME result => SOME (n::result)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2387
                        | NONE        => search (xs, n+1))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2388
                    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2389
                      search (xs, 0)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2390
                    end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2391
                  (* returns the index of the constructor and indices for *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2392
                  (* its arguments that generate the 'elem'-th element of *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2393
                  (* the datatype given by 'idx'                          *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2394
                  (* int -> int -> int * int list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2395
                  fun get_cargs idx elem =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2396
                    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2397
                      (* int * interpretation list -> int * int list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2398
                      fun get_cargs_rec (_, []) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2399
                        raise REFUTE ("IDT_recursion_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2400
                          "no matching constructor found for element "
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2401
                          ^ string_of_int elem ^ " in datatype "
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2402
                          ^ Sign.string_of_typ thy IDT ^ " (datatype index "
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2403
                          ^ string_of_int idx ^ ")")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2404
                        | get_cargs_rec (n, x::xs) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2405
                        (case get_args x elem of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2406
                          SOME args => (n, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2407
                        | NONE      => get_cargs_rec (n+1, xs))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2408
                    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2409
                      get_cargs_rec (0, lookup mc_intrs idx)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2410
                    end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2411
                  (* returns the number of constructors in datatypes that *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2412
                  (* occur in the descriptor 'descr' before the datatype  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2413
                  (* given by 'idx'                                       *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2414
                  fun get_coffset idx =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2415
                    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2416
                      fun get_coffset_acc _ [] =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2417
                        raise REFUTE ("IDT_recursion_interpreter", "index "
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2418
                          ^ string_of_int idx ^ " not found in descriptor")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2419
                        | get_coffset_acc sum ((i, (_, _, cs))::descr') =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2420
                        if i=idx then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2421
                          sum
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2422
                        else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2423
                          get_coffset_acc (sum + length cs) descr'
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2424
                    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2425
                      get_coffset_acc 0 descr
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2426
                    end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2427
                  (* computes one entry in INTRS, and recursively all      *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2428
                  (* entries needed for it, where 'idx' gives the datatype *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2429
                  (* and 'elem' the element of it                          *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2430
                  (* int -> int -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2431
                  fun compute_array_entry idx elem =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2432
                    case Array.sub (lookup INTRS idx, elem) of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2433
                      SOME result =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2434
                      (* simply return the previously computed result *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2435
                      result
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2436
                    | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2437
                      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2438
                        (* int * int list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2439
                        val (c, args) = get_cargs idx elem
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2440
                        (* interpretation * int list -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2441
                        fun select_subtree (tr, []) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2442
                          tr  (* return the whole tree *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2443
                          | select_subtree (Leaf _, _) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2444
                          raise REFUTE ("IDT_recursion_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2445
                            "interpretation for parameter is a leaf; "
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2446
                            ^ "cannot select a subtree")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2447
                          | select_subtree (Node tr, x::xs) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2448
                          select_subtree (List.nth (tr, x), xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2449
                        (* select the correct subtree of the parameter *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2450
                        (* corresponding to constructor 'c'            *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2451
                        val p_intr = select_subtree (List.nth
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2452
                          (p_intrs, get_coffset idx + c), args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2453
                        (* find the indices of the constructor's recursive *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2454
                        (* arguments                                       *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2455
                        val (_, _, constrs) = lookup descr idx
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2456
                        val constr_args     = (snd o List.nth) (constrs, c)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2457
                        val rec_args        = List.filter
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2458
                          (DatatypeAux.is_rec_type o fst) (constr_args ~~ args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2459
                        val rec_args'       = map (fn (dtyp, elem) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2460
                          (DatatypeAux.dest_DtRec dtyp, elem)) rec_args
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2461
                        (* apply 'p_intr' to recursively computed results *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2462
                        val result = foldl (fn ((idx, elem), intr) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2463
                          interpretation_apply (intr,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2464
                          compute_array_entry idx elem)) p_intr rec_args'
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2465
                        (* update 'INTRS' *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2466
                        val _ = Array.update (lookup INTRS idx, elem,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2467
                          SOME result)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2468
                      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2469
                        result
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2470
                      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2471
                  (* compute all entries in INTRS for the current datatype *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2472
                  (* (given by 'index')                                    *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2473
                  (* TODO: we can use Array.modifyi instead once PolyML's *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2474
                  (*       Array signature conforms to the ML standard    *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2475
                  (* (int * 'a -> 'a) -> 'a array -> unit *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2476
                  fun modifyi f arr =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2477
                    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2478
                      val size = Array.length arr
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2479
                      fun modifyi_loop i =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2480
                        if i < size then (
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2481
                          Array.update (arr, i, f (i, Array.sub (arr, i)));
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2482
                          modifyi_loop (i+1)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2483
                        ) else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2484
                          ()
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2485
                    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2486
                      modifyi_loop 0
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2487
                    end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2488
                  val _ = modifyi (fn (i, _) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2489
                    SOME (compute_array_entry index i)) (lookup INTRS index)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2490
                  (* 'a Array.array -> 'a list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2491
                  fun toList arr =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2492
                    Array.foldr op:: [] arr
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2493
                in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2494
                  (* return the part of 'INTRS' that corresponds to the *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2495
                  (* current datatype                                   *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2496
                  SOME ((Node o map Option.valOf o toList o lookup INTRS)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2497
                    index, model', args')
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2498
                end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2499
            end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2500
          else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2501
            NONE  (* not a recursion operator of this datatype *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2502
        ) (DatatypePackage.get_datatypes thy) NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2503
    | _ =>  (* head of term is not a constant *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2504
      NONE;
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  2505
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2506
  (* theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2507
    (interpretation * model * arguments) option *)
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  2508
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2509
  (* only an optimization: 'card' could in principle be interpreted with *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2510
  (* interpreters available already (using its definition), but the code *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2511
  (* below is more efficient                                             *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  2512
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2513
  fun Finite_Set_card_interpreter thy model args t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2514
    case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2515
      Const ("Finite_Set.card",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2516
        Type ("fun", [Type ("set", [T]), Type ("nat", [])])) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2517
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2518
        val (i_nat, _, _) = interpret thy model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2519
          {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2520
          (Free ("dummy", Type ("nat", [])))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2521
        val size_nat      = size_of_type i_nat
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2522
        val (i_set, _, _) = interpret thy model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2523
          {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2524
          (Free ("dummy", Type ("set", [T])))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2525
        val constants     = make_constants i_set
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2526
        (* interpretation -> int *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2527
        fun number_of_elements (Node xs) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2528
          Library.foldl (fn (n, x) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2529
            if x=TT then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2530
              n+1
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2531
            else if x=FF then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2532
              n
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2533
            else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2534
              raise REFUTE ("Finite_Set_card_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2535
                "interpretation for set type does not yield a Boolean"))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2536
            (0, xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2537
          | number_of_elements (Leaf _) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2538
          raise REFUTE ("Finite_Set_card_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2539
            "interpretation for set type is a leaf")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2540
        (* takes an interpretation for a set and returns an interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2541
        (* for a 'nat'                                                     *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2542
        (* interpretation -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2543
        fun card i =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2544
          let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2545
            val n = number_of_elements i
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2546
          in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2547
            if n<size_nat then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2548
              Leaf ((replicate n False) @ True ::
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2549
                (replicate (size_nat-n-1) False))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2550
            else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2551
              Leaf (replicate size_nat False)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2552
          end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2553
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2554
        SOME (Node (map card constants), model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2555
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2556
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2557
      NONE;
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  2558
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2559
  (* theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2560
    (interpretation * model * arguments) option *)
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  2561
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2562
  (* only an optimization: 'Finites' could in principle be interpreted with *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2563
  (* interpreters available already (using its definition), but the code    *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2564
  (* below is more efficient                                                *)
15571
c166086feace interpreter for Finite_Set.Finites added
webertj
parents: 15570
diff changeset
  2565
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2566
  fun Finite_Set_Finites_interpreter thy model args t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2567
    case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2568
      Const ("Finite_Set.Finites", Type ("set", [Type ("set", [T])])) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2569
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2570
        val (i_set, _, _) = interpret thy model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2571
          {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2572
          (Free ("dummy", Type ("set", [T])))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2573
        val size_set      = size_of_type i_set
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2574
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2575
        (* we only consider finite models anyway, hence EVERY set is in *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2576
        (* "Finites"                                                    *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2577
        SOME (Node (replicate size_set TT), model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2578
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2579
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2580
      NONE;
15571
c166086feace interpreter for Finite_Set.Finites added
webertj
parents: 15570
diff changeset
  2581
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2582
  (* theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2583
    (interpretation * model * arguments) option *)
15571
c166086feace interpreter for Finite_Set.Finites added
webertj
parents: 15570
diff changeset
  2584
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2585
  (* only an optimization: 'finite' could in principle be interpreted with  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2586
  (* interpreters available already (using its definition), but the code    *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2587
  (* below is more efficient                                                *)
22093
98e3e9f00192 interpreter for Finite_Set.finite added
webertj
parents: 22092
diff changeset
  2588
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2589
  fun Finite_Set_finite_interpreter thy model args t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2590
    case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2591
      Const ("Finite_Set.finite",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2592
        Type ("fun", [Type ("set", [T]), Type ("bool", [])])) $ _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2593
        (* we only consider finite models anyway, hence EVERY set is *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2594
        (* "finite"                                                  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2595
        SOME (TT, model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2596
    | Const ("Finite_Set.finite",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2597
        Type ("fun", [Type ("set", [T]), Type ("bool", [])])) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2598
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2599
        val (i_set, _, _) = interpret thy model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2600
          {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2601
          (Free ("dummy", Type ("set", [T])))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2602
        val size_set      = size_of_type i_set
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2603
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2604
        (* we only consider finite models anyway, hence EVERY set is *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2605
        (* "finite"                                                  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2606
        SOME (Node (replicate size_set TT), model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2607
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2608
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2609
      NONE;
22093
98e3e9f00192 interpreter for Finite_Set.finite added
webertj
parents: 22092
diff changeset
  2610
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2611
  (* theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2612
    (interpretation * model * arguments) option *)
22093
98e3e9f00192 interpreter for Finite_Set.finite added
webertj
parents: 22092
diff changeset
  2613
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2614
  (* only an optimization: 'Orderings.less' could in principle be            *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2615
  (* interpreted with interpreters available already (using its definition), *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2616
  (* but the code below is more efficient                                    *)
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  2617
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2618
  fun Nat_less_interpreter thy model args t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2619
    case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2620
      Const ("Orderings.less", Type ("fun", [Type ("nat", []),
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2621
        Type ("fun", [Type ("nat", []), Type ("bool", [])])])) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2622
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2623
        val (i_nat, _, _) = interpret thy model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2624
          {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2625
          (Free ("dummy", Type ("nat", [])))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2626
        val size_nat      = size_of_type i_nat
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2627
        (* int -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2628
        (* the 'n'-th nat is not less than the first 'n' nats, while it *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2629
        (* is less than the remaining 'size_nat - n' nats               *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2630
        fun less n = Node ((replicate n FF) @ (replicate (size_nat - n) TT))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2631
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2632
        SOME (Node (map less (1 upto size_nat)), model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2633
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2634
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2635
      NONE;
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  2636
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2637
  (* theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2638
    (interpretation * model * arguments) option *)
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  2639
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2640
  (* only an optimization: 'HOL.plus' could in principle be interpreted with *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2641
  (* interpreters available already (using its definition), but the code     *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2642
  (* below is more efficient                                                 *)
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  2643
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2644
  fun Nat_plus_interpreter thy model args t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2645
    case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2646
      Const ("HOL.plus", Type ("fun", [Type ("nat", []),
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2647
        Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2648
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2649
        val (i_nat, _, _) = interpret thy model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2650
          {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2651
          (Free ("dummy", Type ("nat", [])))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2652
        val size_nat      = size_of_type i_nat
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2653
        (* int -> int -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2654
        fun plus m n =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2655
          let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2656
            val element = (m+n)+1
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2657
          in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2658
            if element > size_nat then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2659
              Leaf (replicate size_nat False)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2660
            else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2661
              Leaf ((replicate (element-1) False) @ True ::
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2662
                (replicate (size_nat - element) False))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2663
          end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2664
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2665
        SOME (Node (map (fn m => Node (map (plus m) (0 upto size_nat-1)))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2666
          (0 upto size_nat-1)), model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2667
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2668
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2669
      NONE;
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  2670
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2671
  (* theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2672
    (interpretation * model * arguments) option *)
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  2673
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2674
  (* only an optimization: 'HOL.minus' could in principle be interpreted *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2675
  (* with interpreters available already (using its definition), but the *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2676
  (* code below is more efficient                                        *)
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  2677
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2678
  fun Nat_minus_interpreter thy model args t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2679
    case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2680
      Const ("HOL.minus", Type ("fun", [Type ("nat", []),
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2681
        Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2682
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2683
        val (i_nat, _, _) = interpret thy model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2684
          {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2685
          (Free ("dummy", Type ("nat", [])))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2686
        val size_nat      = size_of_type i_nat
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2687
        (* int -> int -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2688
        fun minus m n =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2689
          let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2690
            val element = Int.max (m-n, 0) + 1
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2691
          in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2692
            Leaf ((replicate (element-1) False) @ True ::
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2693
              (replicate (size_nat - element) False))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2694
          end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2695
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2696
        SOME (Node (map (fn m => Node (map (minus m) (0 upto size_nat-1)))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2697
          (0 upto size_nat-1)), model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2698
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2699
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2700
      NONE;
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  2701
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2702
  (* theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2703
    (interpretation * model * arguments) option *)
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  2704
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2705
  (* only an optimization: 'HOL.times' could in principle be interpreted with *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2706
  (* interpreters available already (using its definition), but the code      *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2707
  (* below is more efficient                                                  *)
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  2708
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2709
  fun Nat_times_interpreter thy model args t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2710
    case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2711
      Const ("HOL.times", Type ("fun", [Type ("nat", []),
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2712
        Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2713
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2714
        val (i_nat, _, _) = interpret thy model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2715
          {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2716
          (Free ("dummy", Type ("nat", [])))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2717
        val size_nat      = size_of_type i_nat
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2718
        (* nat -> nat -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2719
        fun mult m n =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2720
          let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2721
            val element = (m*n)+1
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2722
          in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2723
            if element > size_nat then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2724
              Leaf (replicate size_nat False)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2725
            else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2726
              Leaf ((replicate (element-1) False) @ True ::
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2727
                (replicate (size_nat - element) False))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2728
          end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2729
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2730
        SOME (Node (map (fn m => Node (map (mult m) (0 upto size_nat-1)))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2731
          (0 upto size_nat-1)), model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2732
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2733
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2734
      NONE;
15547
f08e2d83681e major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
webertj
parents: 15531
diff changeset
  2735
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2736
  (* theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2737
    (interpretation * model * arguments) option *)
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15611
diff changeset
  2738
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2739
  (* only an optimization: 'op @' could in principle be interpreted with *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2740
  (* interpreters available already (using its definition), but the code *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2741
  (* below is more efficient                                             *)
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15611
diff changeset
  2742
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2743
  fun List_append_interpreter thy model args t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2744
    case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2745
      Const ("List.op @", Type ("fun", [Type ("List.list", [T]), Type ("fun",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2746
        [Type ("List.list", [_]), Type ("List.list", [_])])])) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2747
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2748
        val (i_elem, _, _) = interpret thy model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2749
          {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2750
          (Free ("dummy", T))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2751
        val size_elem      = size_of_type i_elem
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2752
        val (i_list, _, _) = interpret thy model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2753
          {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2754
          (Free ("dummy", Type ("List.list", [T])))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2755
        val size_list      = size_of_type i_list
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2756
        (* power (a, b) computes a^b, for a>=0, b>=0 *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2757
        (* int * int -> int *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2758
        fun power (a, 0) = 1
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2759
          | power (a, 1) = a
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2760
          | power (a, b) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2761
          let val ab = power(a, b div 2) in ab * ab * power(a, b mod 2) end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2762
        (* log (a, b) computes floor(log_a(b)), i.e. the largest integer x *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2763
        (* s.t. a^x <= b, for a>=2, b>=1                                   *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2764
        (* int * int -> int *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2765
        fun log (a, b) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2766
          let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2767
            fun logloop (ax, x) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2768
              if ax > b then x-1 else logloop (a * ax, x+1)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2769
          in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2770
            logloop (1, 0)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2771
          end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2772
        (* nat -> nat -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2773
        fun append m n =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2774
          let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2775
            (* The following formula depends on the order in which lists are *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2776
            (* enumerated by the 'IDT_constructor_interpreter'.  It took me  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2777
            (* a little while to come up with this formula.                  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2778
            val element = n + m * (if size_elem = 1 then 1
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2779
              else power (size_elem, log (size_elem, n+1))) + 1
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2780
          in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2781
            if element > size_list then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2782
              Leaf (replicate size_list False)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2783
            else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2784
              Leaf ((replicate (element-1) False) @ True ::
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2785
                (replicate (size_list - element) False))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2786
          end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2787
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2788
        SOME (Node (map (fn m => Node (map (append m) (0 upto size_list-1)))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2789
          (0 upto size_list-1)), model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2790
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2791
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2792
      NONE;
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15611
diff changeset
  2793
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2794
  (* theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2795
    (interpretation * model * arguments) option *)
16050
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15794
diff changeset
  2796
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2797
  (* only an optimization: 'lfp' could in principle be interpreted with  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2798
  (* interpreters available already (using its definition), but the code *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2799
  (* below is more efficient                                             *)
16050
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15794
diff changeset
  2800
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2801
  fun Lfp_lfp_interpreter thy model args t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2802
    case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2803
      Const ("Lfp.lfp", Type ("fun", [Type ("fun",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2804
        [Type ("set", [T]), Type ("set", [_])]), Type ("set", [_])])) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2805
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2806
        val (i_elem, _, _) = interpret thy model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2807
          {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2808
          (Free ("dummy", T))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2809
        val size_elem      = size_of_type i_elem
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2810
        (* the universe (i.e. the set that contains every element) *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2811
        val i_univ         = Node (replicate size_elem TT)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2812
        (* all sets with elements from type 'T' *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2813
        val (i_set, _, _)  = interpret thy model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2814
          {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2815
          (Free ("dummy", Type ("set", [T])))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2816
        val i_sets         = make_constants i_set
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2817
        (* all functions that map sets to sets *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2818
        val (i_fun, _, _)  = interpret thy model {maxvars=0, def_eq=false,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2819
          next_idx=1, bounds=[], wellformed=True} (Free ("dummy",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2820
          Type ("fun", [Type ("set", [T]), Type ("set", [T])])))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2821
        val i_funs         = make_constants i_fun
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2822
        (* "lfp(f) == Inter({u. f(u) <= u})" *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2823
        (* interpretation * interpretation -> bool *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2824
        fun is_subset (Node subs, Node sups) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2825
          List.all (fn (sub, sup) => (sub = FF) orelse (sup = TT))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2826
            (subs ~~ sups)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2827
          | is_subset (_, _) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2828
          raise REFUTE ("Lfp_lfp_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2829
            "is_subset: interpretation for set is not a node")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2830
        (* interpretation * interpretation -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2831
        fun intersection (Node xs, Node ys) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2832
          Node (map (fn (x, y) => if x=TT andalso y=TT then TT else FF)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2833
            (xs ~~ ys))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2834
          | intersection (_, _) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2835
          raise REFUTE ("Lfp_lfp_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2836
            "intersection: interpretation for set is not a node")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2837
        (* interpretation -> interpretaion *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2838
        fun lfp (Node resultsets) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2839
          foldl (fn ((set, resultset), acc) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2840
            if is_subset (resultset, set) then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2841
              intersection (acc, set)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2842
            else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2843
              acc) i_univ (i_sets ~~ resultsets)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2844
          | lfp _ =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2845
            raise REFUTE ("Lfp_lfp_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2846
              "lfp: interpretation for function is not a node")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2847
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2848
        SOME (Node (map lfp i_funs), model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2849
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2850
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2851
      NONE;
16050
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15794
diff changeset
  2852
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2853
  (* theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2854
    (interpretation * model * arguments) option *)
16050
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15794
diff changeset
  2855
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2856
  (* only an optimization: 'gfp' could in principle be interpreted with  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2857
  (* interpreters available already (using its definition), but the code *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2858
  (* below is more efficient                                             *)
16050
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15794
diff changeset
  2859
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2860
  fun Gfp_gfp_interpreter thy model args t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2861
    case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2862
      Const ("Gfp.gfp", Type ("fun", [Type ("fun",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2863
        [Type ("set", [T]), Type ("set", [_])]), Type ("set", [_])])) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2864
      let nonfix union (* because "union" is used below *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2865
        val (i_elem, _, _) = interpret thy model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2866
          {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2867
          (Free ("dummy", T))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2868
        val size_elem      = size_of_type i_elem
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2869
        (* the universe (i.e. the set that contains every element) *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2870
        val i_univ         = Node (replicate size_elem TT)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2871
        (* all sets with elements from type 'T' *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2872
        val (i_set, _, _)  = interpret thy model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2873
          {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2874
          (Free ("dummy", Type ("set", [T])))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2875
        val i_sets         = make_constants i_set
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2876
        (* all functions that map sets to sets *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2877
        val (i_fun, _, _)  = interpret thy model {maxvars=0, def_eq=false,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2878
          next_idx=1, bounds=[], wellformed=True} (Free ("dummy",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2879
          Type ("fun", [Type ("set", [T]), Type ("set", [T])])))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2880
        val i_funs         = make_constants i_fun
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2881
        (* "gfp(f) == Union({u. u <= f(u)})" *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2882
        (* interpretation * interpretation -> bool *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2883
        fun is_subset (Node subs, Node sups) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2884
          List.all (fn (sub, sup) => (sub = FF) orelse (sup = TT))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2885
            (subs ~~ sups)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2886
          | is_subset (_, _) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2887
          raise REFUTE ("Gfp_gfp_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2888
            "is_subset: interpretation for set is not a node")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2889
        (* interpretation * interpretation -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2890
        fun union (Node xs, Node ys) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2891
            Node (map (fn (x,y) => if x=TT orelse y=TT then TT else FF)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2892
                 (xs ~~ ys))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2893
          | union (_, _) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2894
          raise REFUTE ("Gfp_gfp_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2895
            "union: interpretation for set is not a node")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2896
        (* interpretation -> interpretaion *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2897
        fun gfp (Node resultsets) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2898
          foldl (fn ((set, resultset), acc) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2899
            if is_subset (set, resultset) then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2900
              union (acc, set)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2901
            else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2902
              acc) i_univ (i_sets ~~ resultsets)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2903
          | gfp _ =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2904
            raise REFUTE ("Gfp_gfp_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2905
              "gfp: interpretation for function is not a node")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2906
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2907
        SOME (Node (map gfp i_funs), model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2908
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2909
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2910
      NONE;
16050
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15794
diff changeset
  2911
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2912
  (* theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2913
    (interpretation * model * arguments) option *)
21267
5294ecae6708 interpreters for fst and snd added
webertj
parents: 21098
diff changeset
  2914
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2915
  (* only an optimization: 'fst' could in principle be interpreted with  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2916
  (* interpreters available already (using its definition), but the code *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2917
  (* below is more efficient                                             *)
21267
5294ecae6708 interpreters for fst and snd added
webertj
parents: 21098
diff changeset
  2918
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2919
  fun Product_Type_fst_interpreter thy model args t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2920
    case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2921
      Const ("fst", Type ("fun", [Type ("*", [T, U]), _])) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2922
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2923
        val (iT, _, _) = interpret thy model {maxvars=0, def_eq=false,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2924
          next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2925
        val is_T       = make_constants iT
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2926
        val (iU, _, _) = interpret thy model {maxvars=0, def_eq=false,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2927
          next_idx=1, bounds=[], wellformed=True} (Free ("dummy", U))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2928
        val size_U     = size_of_type iU
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2929
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2930
        SOME (Node (List.concat (map (replicate size_U) is_T)), model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2931
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2932
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2933
      NONE;
21267
5294ecae6708 interpreters for fst and snd added
webertj
parents: 21098
diff changeset
  2934
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2935
  (* theory -> model -> arguments -> Term.term ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2936
    (interpretation * model * arguments) option *)
21267
5294ecae6708 interpreters for fst and snd added
webertj
parents: 21098
diff changeset
  2937
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2938
  (* only an optimization: 'snd' could in principle be interpreted with  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2939
  (* interpreters available already (using its definition), but the code *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2940
  (* below is more efficient                                             *)
21267
5294ecae6708 interpreters for fst and snd added
webertj
parents: 21098
diff changeset
  2941
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2942
  fun Product_Type_snd_interpreter thy model args t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2943
    case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2944
      Const ("snd", Type ("fun", [Type ("*", [T, U]), _])) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2945
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2946
        val (iT, _, _) = interpret thy model {maxvars=0, def_eq=false,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2947
          next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2948
        val size_T     = size_of_type iT
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2949
        val (iU, _, _) = interpret thy model {maxvars=0, def_eq=false,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2950
          next_idx=1, bounds=[], wellformed=True} (Free ("dummy", U))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2951
        val is_U       = make_constants iU
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2952
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2953
        SOME (Node (List.concat (replicate size_T is_U)), model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2954
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2955
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2956
      NONE;
21267
5294ecae6708 interpreters for fst and snd added
webertj
parents: 21098
diff changeset
  2957
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  2958
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  2959
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  2960
(* PRINTERS                                                                  *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  2961
(* ------------------------------------------------------------------------- *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  2962
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2963
  (* theory -> model -> Term.term -> interpretation -> (int -> bool) ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2964
    Term.term option *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  2965
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2966
  fun stlc_printer thy model t intr assignment =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2967
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2968
    (* Term.term -> Term.typ option *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2969
    fun typeof (Free (_, T))  = SOME T
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2970
      | typeof (Var (_, T))   = SOME T
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2971
      | typeof (Const (_, T)) = SOME T
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2972
      | typeof _              = NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2973
    (* string -> string *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2974
    fun strip_leading_quote s =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2975
      (implode o (fn [] => [] | x::xs => if x="'" then xs else x::xs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2976
        o explode) s
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2977
    (* Term.typ -> string *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2978
    fun string_of_typ (Type (s, _))     = s
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2979
      | string_of_typ (TFree (x, _))    = strip_leading_quote x
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2980
      | string_of_typ (TVar ((x,i), _)) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2981
      strip_leading_quote x ^ string_of_int i
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2982
    (* interpretation -> int *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2983
    fun index_from_interpretation (Leaf xs) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2984
      find_index (PropLogic.eval assignment) xs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2985
      | index_from_interpretation _ =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2986
      raise REFUTE ("stlc_printer",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2987
        "interpretation for ground type is not a leaf")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2988
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2989
    case typeof t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2990
      SOME T =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2991
      (case T of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2992
        Type ("fun", [T1, T2]) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2993
        let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2994
          (* create all constants of type 'T1' *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2995
          val (i, _, _) = interpret thy model {maxvars=0, def_eq=false,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2996
            next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T1))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2997
          val constants = make_constants i
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2998
          (* interpretation list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2999
          val results = (case intr of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3000
              Node xs => xs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3001
            | _       => raise REFUTE ("stlc_printer",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3002
              "interpretation for function type is a leaf"))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3003
          (* Term.term list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3004
          val pairs = map (fn (arg, result) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3005
            HOLogic.mk_prod
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3006
              (print thy model (Free ("dummy", T1)) arg assignment,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3007
               print thy model (Free ("dummy", T2)) result assignment))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3008
            (constants ~~ results)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3009
          (* Term.typ *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3010
          val HOLogic_prodT = HOLogic.mk_prodT (T1, T2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3011
          val HOLogic_setT  = HOLogic.mk_setT HOLogic_prodT
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3012
          (* Term.term *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3013
          val HOLogic_empty_set = Const ("{}", HOLogic_setT)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3014
          val HOLogic_insert    =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3015
            Const ("insert", HOLogic_prodT --> HOLogic_setT --> HOLogic_setT)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3016
        in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3017
          SOME (foldr (fn (pair, acc) => HOLogic_insert $ pair $ acc)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3018
            HOLogic_empty_set pairs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3019
        end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3020
      | Type ("prop", [])      =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3021
        (case index_from_interpretation intr of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3022
          ~1 => SOME (HOLogic.mk_Trueprop (Const ("arbitrary", HOLogic.boolT)))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3023
        | 0  => SOME (HOLogic.mk_Trueprop HOLogic.true_const)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3024
        | 1  => SOME (HOLogic.mk_Trueprop HOLogic.false_const)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3025
        | _  => raise REFUTE ("stlc_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3026
          "illegal interpretation for a propositional value"))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3027
      | Type _  => if index_from_interpretation intr = (~1) then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3028
          SOME (Const ("arbitrary", T))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3029
        else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3030
          SOME (Const (string_of_typ T ^
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3031
            string_of_int (index_from_interpretation intr), T))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3032
      | TFree _ => if index_from_interpretation intr = (~1) then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3033
          SOME (Const ("arbitrary", T))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3034
        else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3035
          SOME (Const (string_of_typ T ^
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3036
            string_of_int (index_from_interpretation intr), T))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3037
      | TVar _  => if index_from_interpretation intr = (~1) then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3038
          SOME (Const ("arbitrary", T))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3039
        else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3040
          SOME (Const (string_of_typ T ^
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3041
            string_of_int (index_from_interpretation intr), T)))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3042
    | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3043
      NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3044
  end;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3045
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3046
  (* theory -> model -> Term.term -> interpretation -> (int -> bool) ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3047
    string option *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3048
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3049
  fun set_printer thy model t intr assignment =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3050
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3051
    (* Term.term -> Term.typ option *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3052
    fun typeof (Free (_, T))  = SOME T
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3053
      | typeof (Var (_, T))   = SOME T
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3054
      | typeof (Const (_, T)) = SOME T
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3055
      | typeof _              = NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3056
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3057
    case typeof t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3058
      SOME (Type ("set", [T])) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3059
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3060
        (* create all constants of type 'T' *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3061
        val (i, _, _) = interpret thy model {maxvars=0, def_eq=false,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3062
          next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3063
        val constants = make_constants i
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3064
        (* interpretation list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3065
        val results = (case intr of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3066
            Node xs => xs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3067
          | _       => raise REFUTE ("set_printer",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3068
            "interpretation for set type is a leaf"))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3069
        (* Term.term list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3070
        val elements = List.mapPartial (fn (arg, result) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3071
          case result of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3072
            Leaf [fmTrue, fmFalse] =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3073
            if PropLogic.eval assignment fmTrue then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3074
              SOME (print thy model (Free ("dummy", T)) arg assignment)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3075
            else (* if PropLogic.eval assignment fmFalse then *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3076
              NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3077
          | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3078
            raise REFUTE ("set_printer",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3079
              "illegal interpretation for a Boolean value"))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3080
          (constants ~~ results)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3081
        (* Term.typ *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3082
        val HOLogic_setT  = HOLogic.mk_setT T
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3083
        (* Term.term *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3084
        val HOLogic_empty_set = Const ("{}", HOLogic_setT)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3085
        val HOLogic_insert    =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3086
          Const ("insert", T --> HOLogic_setT --> HOLogic_setT)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3087
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3088
        SOME (Library.foldl (fn (acc, elem) => HOLogic_insert $ elem $ acc)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3089
          (HOLogic_empty_set, elements))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3090
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3091
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3092
      NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3093
  end;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3094
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3095
  (* theory -> model -> Term.term -> interpretation -> (int -> bool) ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3096
    Term.term option *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3097
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3098
  fun IDT_printer thy model t intr assignment =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3099
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3100
    (* Term.term -> Term.typ option *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3101
    fun typeof (Free (_, T))  = SOME T
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3102
      | typeof (Var (_, T))   = SOME T
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3103
      | typeof (Const (_, T)) = SOME T
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3104
      | typeof _              = NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3105
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3106
    case typeof t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3107
      SOME (Type (s, Ts)) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3108
      (case DatatypePackage.get_datatype thy s of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3109
        SOME info =>  (* inductive datatype *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3110
        let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3111
          val (typs, _)           = model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3112
          val index               = #index info
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3113
          val descr               = #descr info
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3114
          val (_, dtyps, constrs) = lookup descr index
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3115
          val typ_assoc           = dtyps ~~ Ts
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3116
          (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3117
          val _ = (if Library.exists (fn d =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3118
              case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3119
            then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3120
              raise REFUTE ("IDT_printer", "datatype argument (for type " ^
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3121
                Sign.string_of_typ thy (Type (s, Ts)) ^ ") is not a variable")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3122
            else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3123
              ())
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3124
          (* the index of the element in the datatype *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3125
          val element = (case intr of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3126
              Leaf xs => find_index (PropLogic.eval assignment) xs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3127
            | Node _  => raise REFUTE ("IDT_printer",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3128
              "interpretation is not a leaf"))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3129
        in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3130
          if element < 0 then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3131
            SOME (Const ("arbitrary", Type (s, Ts)))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3132
          else let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3133
            (* takes a datatype constructor, and if for some arguments this  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3134
            (* constructor generates the datatype's element that is given by *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3135
            (* 'element', returns the constructor (as a term) as well as the *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3136
            (* indices of the arguments                                      *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3137
            (* string * DatatypeAux.dtyp list ->
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3138
              (Term.term * int list) option *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3139
            fun get_constr_args (cname, cargs) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3140
              let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3141
                val cTerm      = Const (cname,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3142
                  map (typ_of_dtyp descr typ_assoc) cargs ---> Type (s, Ts))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3143
                val (iC, _, _) = interpret thy (typs, []) {maxvars=0,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3144
                  def_eq=false, next_idx=1, bounds=[], wellformed=True} cTerm
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3145
                (* interpretation -> int list option *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3146
                fun get_args (Leaf xs) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3147
                  if find_index_eq True xs = element then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3148
                    SOME []
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3149
                  else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3150
                    NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3151
                  | get_args (Node xs) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3152
                  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3153
                    (* interpretation * int -> int list option *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3154
                    fun search ([], _) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3155
                      NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3156
                      | search (x::xs, n) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3157
                      (case get_args x of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3158
                        SOME result => SOME (n::result)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3159
                      | NONE        => search (xs, n+1))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3160
                  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3161
                    search (xs, 0)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3162
                  end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3163
              in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3164
                Option.map (fn args => (cTerm, cargs, args)) (get_args iC)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3165
              end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3166
            (* Term.term * DatatypeAux.dtyp list * int list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3167
            val (cTerm, cargs, args) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3168
              (case get_first get_constr_args constrs of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3169
                SOME x => x
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3170
              | NONE   => raise REFUTE ("IDT_printer",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3171
                "no matching constructor found for element " ^
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3172
                string_of_int element))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3173
            val argsTerms = map (fn (d, n) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3174
              let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3175
                val dT        = typ_of_dtyp descr typ_assoc d
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3176
                val (i, _, _) = interpret thy (typs, []) {maxvars=0,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3177
                  def_eq=false, next_idx=1, bounds=[], wellformed=True}
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3178
                  (Free ("dummy", dT))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3179
                (* we only need the n-th element of this list, so there   *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3180
                (* might be a more efficient implementation that does not *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3181
                (* generate all constants                                 *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3182
                val consts    = make_constants i
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3183
              in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3184
                print thy (typs, []) (Free ("dummy", dT))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3185
                  (List.nth (consts, n)) assignment
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3186
              end) (cargs ~~ args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3187
          in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3188
            SOME (Library.foldl op$ (cTerm, argsTerms))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3189
          end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3190
        end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3191
      | NONE =>  (* not an inductive datatype *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3192
        NONE)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3193
    | _ =>  (* a (free or schematic) type variable *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3194
      NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3195
  end;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3196
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3197
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3198
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  3199
(* use 'setup Refute.setup' in an Isabelle theory to initialize the 'Refute' *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  3200
(* structure                                                                 *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3201
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3202
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3203
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  3204
(* Note: the interpreters and printers are used in reverse order; however,   *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  3205
(*       an interpreter that can handle non-atomic terms ends up being       *)
14807
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  3206
(*       applied before the 'stlc_interpreter' breaks the term apart into    *)
e8ccb13d7774 major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
webertj
parents: 14604
diff changeset
  3207
(*       subterms that are then passed to other interpreters!                *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3208
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3209
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3210
  (* (theory -> theory) list *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3211
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3212
  val setup =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3213
     RefuteData.init #>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3214
     add_interpreter "stlc"    stlc_interpreter #>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3215
     add_interpreter "Pure"    Pure_interpreter #>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3216
     add_interpreter "HOLogic" HOLogic_interpreter #>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3217
     add_interpreter "set"     set_interpreter #>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3218
     add_interpreter "IDT"             IDT_interpreter #>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3219
     add_interpreter "IDT_constructor" IDT_constructor_interpreter #>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3220
     add_interpreter "IDT_recursion"   IDT_recursion_interpreter #>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3221
     add_interpreter "Finite_Set.card"    Finite_Set_card_interpreter #>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3222
     add_interpreter "Finite_Set.Finites" Finite_Set_Finites_interpreter #>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3223
     add_interpreter "Finite_Set.finite"  Finite_Set_finite_interpreter #>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3224
     add_interpreter "Nat_Orderings.less" Nat_less_interpreter #>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3225
     add_interpreter "Nat_HOL.plus"       Nat_plus_interpreter #>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3226
     add_interpreter "Nat_HOL.minus"      Nat_minus_interpreter #>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3227
     add_interpreter "Nat_HOL.times"      Nat_times_interpreter #>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3228
     add_interpreter "List.op @" List_append_interpreter #>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3229
     add_interpreter "Lfp.lfp" Lfp_lfp_interpreter #>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3230
     add_interpreter "Gfp.gfp" Gfp_gfp_interpreter #>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3231
     add_interpreter "fst" Product_Type_fst_interpreter #>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3232
     add_interpreter "snd" Product_Type_snd_interpreter #>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3233
     add_printer "stlc" stlc_printer #>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3234
     add_printer "set"  set_printer #>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3235
     add_printer "IDT"  IDT_printer;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3236
22092
ab3dfcef6489 reformatted to 80 chars/line
webertj
parents: 22055
diff changeset
  3237
end  (* structure Refute *)