src/HOL/Tools/refute.ML
author blanchet
Thu, 29 Apr 2010 01:17:14 +0200
changeset 36555 8ff45c2076da
parent 36374 19c0c4b8b445
child 36692 54b64d4ad524
permissions -rw-r--r--
expand combinators in Isar proofs constructed by Sledgehammer; this requires shuffling around a couple of functions previously defined in Refute
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
29265
5b4247055bd7 moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents: 29004
diff changeset
     2
    Author:     Tjark Weber, TU Muenchen
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
     3
14965
7155b319eafa new SAT solver interface
webertj
parents: 14818
diff changeset
     4
Finite model generation for HOL formulas, using a SAT solver.
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
     5
*)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
     6
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
     7
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
     8
(* 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
     9
(* Documentation is available in the Isabelle/Isar theory 'HOL/Refute.thy'.  *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    10
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    11
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    12
signature REFUTE =
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    13
sig
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    14
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    15
  exception REFUTE of string * string
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    16
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    17
(* ------------------------------------------------------------------------- *)
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
    18
(* 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
    19
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    20
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    21
  type params
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    22
  type interpretation
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    23
  type model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    24
  type arguments
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    25
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    26
  exception MAXVARS_EXCEEDED
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    27
33243
17014b1b9353 normalized basic type abbreviations;
wenzelm
parents: 33063
diff changeset
    28
  val add_interpreter : string -> (theory -> model -> arguments -> term ->
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    29
    (interpretation * model * arguments) option) -> theory -> theory
33243
17014b1b9353 normalized basic type abbreviations;
wenzelm
parents: 33063
diff changeset
    30
  val add_printer     : string -> (theory -> model -> typ ->
17014b1b9353 normalized basic type abbreviations;
wenzelm
parents: 33063
diff changeset
    31
    interpretation -> (int -> bool) -> term option) -> theory -> theory
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    32
33243
17014b1b9353 normalized basic type abbreviations;
wenzelm
parents: 33063
diff changeset
    33
  val interpret : theory -> model -> arguments -> term ->
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    34
    (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
    35
33243
17014b1b9353 normalized basic type abbreviations;
wenzelm
parents: 33063
diff changeset
    36
  val print       : theory -> model -> typ -> interpretation -> (int -> bool) -> term
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    37
  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
    38
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    39
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    40
(* Interface                                                                 *)
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
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    43
  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
    44
  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
    45
  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
    46
  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
    47
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
    48
  val find_model : theory -> params -> term list -> term -> bool -> unit
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    49
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    50
  (* tries to find a model for a formula: *)
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
    51
  val satisfy_term :
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
    52
    theory -> (string * string) list -> term list -> term -> unit
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 that refutes a formula: *)
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
    54
  val refute_term :
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
    55
    theory -> (string * string) list -> term list -> term -> unit
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
    56
  val refute_goal :
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
    57
    Proof.context -> (string * string) list -> thm -> int -> unit
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    58
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    59
  val setup : theory -> theory
22092
ab3dfcef6489 reformatted to 80 chars/line
webertj
parents: 22055
diff changeset
    60
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
    61
(* ------------------------------------------------------------------------- *)
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
    62
(* Additional functions used by Nitpick (to be factored out)                 *)
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
    63
(* ------------------------------------------------------------------------- *)
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
    64
33243
17014b1b9353 normalized basic type abbreviations;
wenzelm
parents: 33063
diff changeset
    65
  val close_form : term -> term
17014b1b9353 normalized basic type abbreviations;
wenzelm
parents: 33063
diff changeset
    66
  val get_classdef : theory -> string -> (string * term) option
17014b1b9353 normalized basic type abbreviations;
wenzelm
parents: 33063
diff changeset
    67
  val norm_rhs : term -> term
17014b1b9353 normalized basic type abbreviations;
wenzelm
parents: 33063
diff changeset
    68
  val get_def : theory -> string * typ -> (string * term) option
17014b1b9353 normalized basic type abbreviations;
wenzelm
parents: 33063
diff changeset
    69
  val get_typedef : theory -> typ -> (string * term) option
17014b1b9353 normalized basic type abbreviations;
wenzelm
parents: 33063
diff changeset
    70
  val is_IDT_constructor : theory -> string * typ -> bool
17014b1b9353 normalized basic type abbreviations;
wenzelm
parents: 33063
diff changeset
    71
  val is_IDT_recursor : theory -> string * typ -> bool
17014b1b9353 normalized basic type abbreviations;
wenzelm
parents: 33063
diff changeset
    72
  val is_const_of_class: theory -> string * typ -> bool
17014b1b9353 normalized basic type abbreviations;
wenzelm
parents: 33063
diff changeset
    73
  val string_of_typ : typ -> string
17014b1b9353 normalized basic type abbreviations;
wenzelm
parents: 33063
diff changeset
    74
  val typ_of_dtyp : Datatype.descr -> (Datatype.dtyp * typ) list -> Datatype.dtyp -> typ
22092
ab3dfcef6489 reformatted to 80 chars/line
webertj
parents: 22055
diff changeset
    75
end;  (* signature REFUTE *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    76
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    77
structure Refute : REFUTE =
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    78
struct
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    79
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    80
  open PropLogic;
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
    81
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    82
  (* 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
    83
  (* 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
    84
  (* 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
    85
  (* 'error'.                                                          *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
    86
  exception REFUTE of string * string;  (* ("in function", "cause") *)
14350
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
  (* 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
    89
  (* 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
    90
  exception MAXVARS_EXCEEDED;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    91
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    92
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    93
(* TREES                                                                     *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    94
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    95
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    96
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    97
(* tree: implements an arbitrarily (but finitely) branching tree as a list   *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    98
(*       of (lists of ...) elements                                          *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
    99
(* ------------------------------------------------------------------------- *)
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
  datatype 'a tree =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   102
      Leaf of 'a
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   103
    | Node of ('a tree) list;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   104
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   105
  (* ('a -> 'b) -> 'a tree -> 'b tree *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   106
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   107
  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
   108
    case tr of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   109
      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
   110
    | Node xs => Node (map (tree_map f) xs);
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   111
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   112
  (* ('a * 'b -> 'a) -> 'a * ('b tree) -> 'a *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   113
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   114
  fun tree_foldl f =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   115
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   116
    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
   117
      | 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
   118
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   119
    itl
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   120
  end;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   121
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   122
  (* 'a tree * 'b tree -> ('a * 'b) tree *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   123
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   124
  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
   125
    case t1 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   126
      Leaf x =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   127
      (case t2 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   128
          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
   129
        | 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
   130
            "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
   131
    | Node xs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   132
      (case t2 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   133
          (* '~~' 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
   134
          (* 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
   135
          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
   136
        | 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
   137
            "trees are of different height (first tree is higher)"));
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   138
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   139
(* ------------------------------------------------------------------------- *)
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
   140
(* 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
   141
(*         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
   142
(*                                                                           *)
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
(* The following parameters are supported (and required (!), except for      *)
30314
853778f6ef7d Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
blanchet
parents: 30312
diff changeset
   144
(* "sizes" and "expect"):                                                    *)
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
   145
(*                                                                           *)
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
   146
(* 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
   147
(*                                                                           *)
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
   148
(* "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
   149
(*                       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
   150
(* "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
   151
(* "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
   152
(* "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
   153
(*                       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
   154
(*                       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
   155
(* "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
   156
(* "satsolver"   string  SAT solver to be used.                              *)
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
   157
(* "no_assms"    bool    If "true", assumptions in structured proofs are     *)
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
   158
(*                       not considered.                                     *)
30314
853778f6ef7d Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
blanchet
parents: 30312
diff changeset
   159
(* "expect"      string  Expected result ("genuine", "potential", "none", or *)
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
   160
(*                       "unknown").                                         *)
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
   161
(* ------------------------------------------------------------------------- *)
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
   162
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   163
  type params =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   164
    {
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   165
      sizes    : (string * int) list,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   166
      minsize  : int,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   167
      maxsize  : int,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   168
      maxvars  : int,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   169
      maxtime  : int,
30314
853778f6ef7d Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
blanchet
parents: 30312
diff changeset
   170
      satsolver: string,
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
   171
      no_assms : bool,
30314
853778f6ef7d Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
blanchet
parents: 30312
diff changeset
   172
      expect   : string
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   173
    };
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
   174
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
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   176
(* 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
   177
(*                 'interpretation'                                          *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   178
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   179
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   180
  type interpretation =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   181
    prop_formula list tree;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   182
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   183
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   184
(* 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
   185
(*        terms                                                              *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   186
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   187
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   188
  type model =
33243
17014b1b9353 normalized basic type abbreviations;
wenzelm
parents: 33063
diff changeset
   189
    (typ * int) list * (term * interpretation) list;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   190
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   191
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   192
(* arguments: additional arguments required during interpretation of terms   *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   193
(* ------------------------------------------------------------------------- *)
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
   194
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   195
  type arguments =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   196
    {
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   197
      (* 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
   198
      maxvars   : int,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   199
      (* 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
   200
      def_eq    : bool,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   201
      (* 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
   202
      next_idx  : int,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   203
      bounds    : interpretation list,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   204
      wellformed: prop_formula
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   205
    };
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   206
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   207
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33339
diff changeset
   208
  structure RefuteData = Theory_Data
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22580
diff changeset
   209
  (
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   210
    type T =
33243
17014b1b9353 normalized basic type abbreviations;
wenzelm
parents: 33063
diff changeset
   211
      {interpreters: (string * (theory -> model -> arguments -> term ->
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   212
        (interpretation * model * arguments) option)) list,
33243
17014b1b9353 normalized basic type abbreviations;
wenzelm
parents: 33063
diff changeset
   213
       printers: (string * (theory -> model -> typ -> interpretation ->
17014b1b9353 normalized basic type abbreviations;
wenzelm
parents: 33063
diff changeset
   214
        (int -> bool) -> term option)) list,
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   215
       parameters: string Symtab.table};
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   216
    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
   217
    val extend = I;
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33339
diff changeset
   218
    fun merge
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   219
      ({interpreters = in1, printers = pr1, parameters = pa1},
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33339
diff changeset
   220
       {interpreters = in2, printers = pr2, parameters = pa2}) : T =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   221
      {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
   222
       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
   223
       parameters = Symtab.merge (op=) (pa1, pa2)};
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22580
diff changeset
   224
  );
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   225
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   226
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   227
(* ------------------------------------------------------------------------- *)
15334
d5a92997dc1b exception CANNOT_INTERPRET removed (not needed anymore since the stlc_interpreter can interpret any term)
webertj
parents: 15333
diff changeset
   228
(* 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
   229
(*            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
   230
(*            track of the interpretation of subterms                        *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   231
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   232
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   233
  (* 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
   234
    (interpretation * model * arguments) *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   235
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   236
  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
   237
    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
   238
      (#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
   239
      NONE   => raise REFUTE ("interpret",
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26931
diff changeset
   240
        "no interpreter for term " ^ quote (Syntax.string_of_term_global thy t))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   241
    | SOME x => x;
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   242
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   243
(* ------------------------------------------------------------------------- *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   244
(* print: converts the interpretation 'intr', which must denote a term of    *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   245
(*        type 'T', into a term using a suitable printer                     *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   246
(* ------------------------------------------------------------------------- *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   247
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   248
  (* theory -> model -> Term.typ -> interpretation -> (int -> bool) ->
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   249
    Term.term *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   250
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   251
  fun print thy model T intr assignment =
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   252
    case get_first (fn (_, f) => f thy model T intr assignment)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   253
      (#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
   254
      NONE   => raise REFUTE ("print",
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26931
diff changeset
   255
        "no printer for type " ^ quote (Syntax.string_of_typ_global thy T))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   256
    | SOME x => x;
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   257
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   258
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   259
(* 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
   260
(*              (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
   261
(*              printers                                                     *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   262
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   263
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   264
  (* 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
   265
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   266
  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
   267
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   268
    val (typs, terms) = model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   269
    val typs_msg =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   270
      if null typs then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   271
        "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
   272
      else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   273
        "Size of types: " ^ commas (map (fn (T, i) =>
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26931
diff changeset
   274
          Syntax.string_of_typ_global thy T ^ ": " ^ string_of_int i) typs) ^ "\n"
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   275
    val show_consts_msg =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   276
      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
   277
        "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
   278
      else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   279
        ""
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   280
    val terms_msg =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   281
      if null terms then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   282
        "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
   283
      else
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32950
diff changeset
   284
        cat_lines (map_filter (fn (t, intr) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   285
          (* 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
   286
          if (!show_consts) orelse not (is_Const t) then
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26931
diff changeset
   287
            SOME (Syntax.string_of_term_global thy t ^ ": " ^
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26931
diff changeset
   288
              Syntax.string_of_term_global thy
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   289
                (print thy model (Term.type_of t) intr assignment))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   290
          else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   291
            NONE) terms) ^ "\n"
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   292
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   293
    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
   294
  end;
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   295
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   296
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   297
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   298
(* PARAMETER MANAGEMENT                                                      *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   299
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   300
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   301
  (* 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
   302
    (interpretation * model * arguments) option) -> theory -> theory *)
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
  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
   305
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   306
    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
   307
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   308
    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
   309
      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
   310
      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
   311
    | 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
   312
  end;
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   313
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   314
  (* string -> (theory -> model -> Term.typ -> interpretation ->
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   315
    (int -> bool) -> Term.term option) -> theory -> theory *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   316
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   317
  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
   318
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   319
    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
   320
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   321
    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
   322
      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
   323
      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
   324
    | 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
   325
  end;
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   326
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   327
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   328
(* 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
   329
(*                    parameter table                                        *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   330
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   331
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   332
  (* (string * string) -> theory -> theory *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   333
29004
a5a91f387791 removed Table.extend, NameSpace.extend_table
haftmann
parents: 28524
diff changeset
   334
  fun set_default_param (name, value) = RefuteData.map 
a5a91f387791 removed Table.extend, NameSpace.extend_table
haftmann
parents: 28524
diff changeset
   335
    (fn {interpreters, printers, parameters} =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   336
      {interpreters = interpreters, printers = printers,
29004
a5a91f387791 removed Table.extend, NameSpace.extend_table
haftmann
parents: 28524
diff changeset
   337
        parameters = Symtab.update (name, value) parameters});
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   338
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
   339
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   340
(* 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
   341
(*                    RefuteData's parameter table                           *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   342
(* ------------------------------------------------------------------------- *)
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
  (* theory -> string -> string option *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   345
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   346
  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
   347
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   348
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   349
(* 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
   350
(*                     stored in RefuteData's parameter table                *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   351
(* ------------------------------------------------------------------------- *)
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
  (* theory -> (string * string) list *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   354
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   355
  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
   356
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   357
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   358
(* 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
   359
(*      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
   360
(*      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
   361
(* ------------------------------------------------------------------------- *)
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
  (* theory -> (string * string) list -> params *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   364
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   365
  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
   366
  let
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
   367
    (* (string * string) list * string -> bool *)
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
   368
    fun read_bool (parms, name) =
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
   369
      case AList.lookup (op =) parms name of
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
   370
        SOME "true" => true
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
   371
      | SOME "false" => false
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
   372
      | SOME s => error ("parameter " ^ quote name ^
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
   373
        " (value is " ^ quote s ^ ") must be \"true\" or \"false\"")
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
   374
      | NONE   => error ("parameter " ^ quote name ^
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
   375
          " must be assigned a value")
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   376
    (* (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
   377
    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
   378
      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
   379
        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
   380
          SOME i => i
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   381
        | 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
   382
          " (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
   383
      | 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
   384
          " 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
   385
    (* (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
   386
    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
   387
      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
   388
        SOME s => s
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   389
      | 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
   390
        " 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
   391
    (* 'override' first, defaults last: *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   392
    (* (string * string) list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   393
    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
   394
    (* int *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   395
    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
   396
    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
   397
    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
   398
    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
   399
    (* string *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   400
    val satsolver = read_string (allparams, "satsolver")
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
   401
    val no_assms = read_bool (allparams, "no_assms")
30314
853778f6ef7d Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
blanchet
parents: 30312
diff changeset
   402
    val expect = the_default "" (AList.lookup (op =) allparams "expect")
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   403
    (* 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
   404
    (* 'sizes'                                                            *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   405
    (* 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
   406
    (*       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
   407
    (* (string * int) list *)
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32950
diff changeset
   408
    val sizes     = map_filter
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   409
      (fn (name, value) => Option.map (pair name) (Int.fromString value))
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 33246
diff changeset
   410
      (filter (fn (name, _) => name<>"minsize" andalso name<>"maxsize"
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   411
        andalso name<>"maxvars" andalso name<>"maxtime"
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
   412
        andalso name<>"satsolver" andalso name<>"no_assms") allparams)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   413
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   414
    {sizes=sizes, minsize=minsize, maxsize=maxsize, maxvars=maxvars,
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
   415
      maxtime=maxtime, satsolver=satsolver, no_assms=no_assms, expect=expect}
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   416
  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
   417
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
   418
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
   419
(* ------------------------------------------------------------------------- *)
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
   420
(* 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
   421
(* ------------------------------------------------------------------------- *)
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
   422
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
   423
  fun typ_of_dtyp descr typ_assoc (Datatype_Aux.DtTFree a) =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   424
    (* replace a 'DtTFree' variable by the associated type *)
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
   425
    the (AList.lookup (op =) typ_assoc (Datatype_Aux.DtTFree a))
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
   426
    | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtType (s, ds)) =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   427
    Type (s, map (typ_of_dtyp descr typ_assoc) ds)
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
   428
    | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtRec i) =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   429
    let
29288
253bcf2a5854 avoid polymorphic equality;
wenzelm
parents: 29272
diff changeset
   430
      val (s, ds, _) = the (AList.lookup (op =) descr i)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   431
    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   432
      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
   433
    end;
15335
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
   434
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
   435
(* ------------------------------------------------------------------------- *)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   436
(* 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
   437
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   438
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   439
  (* Term.term -> Term.term *)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   440
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   441
  fun close_form t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   442
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   443
    (* (Term.indexname * Term.typ) list *)
29265
5b4247055bd7 moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents: 29004
diff changeset
   444
    val vars = sort_wrt (fst o fst) (map dest_Var (OldTerm.term_vars t))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   445
  in
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   446
    fold (fn ((x, i), T) => fn t' =>
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   447
      Term.all T $ Abs (x, T, abstract_over (Var ((x, i), T), t'))) vars t
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   448
  end;
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   449
36555
8ff45c2076da expand combinators in Isar proofs constructed by Sledgehammer;
blanchet
parents: 36374
diff changeset
   450
val monomorphic_term = Sledgehammer_Util.monomorphic_term
8ff45c2076da expand combinators in Isar proofs constructed by Sledgehammer;
blanchet
parents: 36374
diff changeset
   451
val specialize_type = Sledgehammer_Util.specialize_type
21985
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
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   454
(* 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
   455
(*                    denotes membership to an axiomatic type class          *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   456
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   457
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   458
  (* theory -> string * Term.typ -> bool *)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   459
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   460
  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
   461
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   462
    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
   463
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   464
    (* 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
   465
    (* 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
   466
    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
   467
  end;
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   468
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   469
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   470
(* 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
   471
(*                     of an inductive datatype in 'thy'                     *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   472
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   473
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   474
  (* theory -> string * Term.typ -> bool *)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   475
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   476
  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
   477
    (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
   478
      Type (s', _) =>
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31737
diff changeset
   479
      (case Datatype.get_constrs thy s' of
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   480
        SOME constrs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   481
        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
   482
          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
   483
      | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   484
        false)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   485
    | _  =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   486
      false);
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   487
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   488
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   489
(* 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
   490
(*                  operator of an inductive datatype in 'thy'               *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   491
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   492
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   493
  (* theory -> string * Term.typ -> bool *)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   494
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   495
  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
   496
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   497
    val rec_names = Symtab.fold (append o #rec_names o snd)
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31737
diff changeset
   498
      (Datatype.get_all thy) []
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   499
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   500
    (* 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
   501
    (* 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
   502
    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
   503
  end;
21985
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
(* ------------------------------------------------------------------------- *)
30275
381ce8d88cb8 Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
blanchet
parents: 30242
diff changeset
   506
(* norm_rhs: maps  f ?t1 ... ?tn == rhs  to  %t1...tn. rhs                   *)
381ce8d88cb8 Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
blanchet
parents: 30242
diff changeset
   507
(* ------------------------------------------------------------------------- *)
381ce8d88cb8 Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
blanchet
parents: 30242
diff changeset
   508
381ce8d88cb8 Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
blanchet
parents: 30242
diff changeset
   509
  fun norm_rhs eqn =
381ce8d88cb8 Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
blanchet
parents: 30242
diff changeset
   510
  let
381ce8d88cb8 Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
blanchet
parents: 30242
diff changeset
   511
    fun lambda (v as Var ((x, _), T)) t = Abs (x, T, abstract_over (v, t))
381ce8d88cb8 Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
blanchet
parents: 30242
diff changeset
   512
      | lambda v t                      = raise TERM ("lambda", [v, t])
381ce8d88cb8 Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
blanchet
parents: 30242
diff changeset
   513
    val (lhs, rhs) = Logic.dest_equals eqn
381ce8d88cb8 Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
blanchet
parents: 30242
diff changeset
   514
    val (_, args)  = Term.strip_comb lhs
381ce8d88cb8 Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
blanchet
parents: 30242
diff changeset
   515
  in
381ce8d88cb8 Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
blanchet
parents: 30242
diff changeset
   516
    fold lambda (rev args) rhs
381ce8d88cb8 Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
blanchet
parents: 30242
diff changeset
   517
  end
381ce8d88cb8 Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
blanchet
parents: 30242
diff changeset
   518
381ce8d88cb8 Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
blanchet
parents: 30242
diff changeset
   519
(* ------------------------------------------------------------------------- *)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   520
(* 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
   521
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   522
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   523
  (* 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
   524
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   525
  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
   526
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   527
    (* (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
   528
    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
   529
      | 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
   530
      (let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   531
        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
   532
        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
   533
        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
   534
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   535
        if s=s' then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   536
          let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   537
            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
   538
            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
   539
            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
   540
          in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   541
            SOME (axname, rhs)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   542
          end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   543
        else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   544
          get_def_ax axioms
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   545
      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
   546
               | 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
   547
               | 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
   548
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   549
    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
   550
  end;
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   551
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   552
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   553
(* 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
   554
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   555
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   556
  (* theory -> Term.typ -> (string * Term.term) option *)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   557
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   558
  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
   559
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   560
    (* (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
   561
    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
   562
      | 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
   563
      (let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   564
        (* 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
   565
        fun type_of_type_definition (Const (s', T')) =
35746
9c97d4e2450e more antiquotations;
wenzelm
parents: 35267
diff changeset
   566
          if s'= @{const_name type_definition} then
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   567
            SOME T'
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   568
          else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   569
            NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   570
          | 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
   571
          | 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
   572
          | 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
   573
          | 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
   574
          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
   575
          | 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
   576
          (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
   577
            SOME x => SOME x
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   578
          | 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
   579
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   580
        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
   581
          SOME T' =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   582
          let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   583
            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
   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
          in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   586
            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
   587
          end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   588
        | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   589
          get_typedef_ax axioms
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   590
      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
   591
               | 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
   592
               | 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
   593
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   594
    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
   595
  end;
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   596
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   597
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   598
(* 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
   599
(*               created by the "axclass" command                            *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   600
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   601
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   602
  (* theory -> string -> (string * Term.term) option *)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   603
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   604
  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
   605
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   606
    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
   607
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   608
    Option.map (pair axname)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   609
      (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
   610
  end;
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   611
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   612
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   613
(* 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
   614
(*              normalizes the result term; certain constants are not        *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   615
(*              unfolded (cf. 'collect_axioms' and the various interpreters  *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   616
(*              below): if the interpretation respects a definition anyway,  *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   617
(*              that definition does not need to be unfolded                 *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   618
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   619
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   620
  (* theory -> Term.term -> Term.term *)
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   621
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   622
  (* 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
   623
  (*       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
   624
  (*       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
   625
  (*       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
   626
  (*       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
   627
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   628
  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
   629
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   630
    (* Term.term -> Term.term *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   631
    fun unfold_loop t =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   632
      case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   633
      (* Pure *)
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   634
        Const (@{const_name all}, _) => t
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   635
      | Const (@{const_name "=="}, _) => t
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   636
      | Const (@{const_name "==>"}, _) => t
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   637
      | Const (@{const_name TYPE}, _) => t  (* axiomatic type classes *)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   638
      (* HOL *)
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   639
      | Const (@{const_name Trueprop}, _) => t
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   640
      | Const (@{const_name Not}, _) => t
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   641
      | (* redundant, since 'True' is also an IDT constructor *)
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   642
        Const (@{const_name True}, _) => t
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   643
      | (* redundant, since 'False' is also an IDT constructor *)
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   644
        Const (@{const_name False}, _) => t
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   645
      | Const (@{const_name undefined}, _) => t
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   646
      | Const (@{const_name The}, _) => t
29820
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
   647
      | Const (@{const_name Hilbert_Choice.Eps}, _) => t
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   648
      | Const (@{const_name All}, _) => t
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   649
      | Const (@{const_name Ex}, _) => t
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   650
      | Const (@{const_name "op ="}, _) => t
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   651
      | Const (@{const_name "op &"}, _) => t
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   652
      | Const (@{const_name "op |"}, _) => t
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   653
      | Const (@{const_name "op -->"}, _) => t
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   654
      (* sets *)
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   655
      | Const (@{const_name Collect}, _) => t
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   656
      | Const (@{const_name "op :"}, _) => t
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   657
      (* other optimizations *)
29820
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
   658
      | Const (@{const_name Finite_Set.card}, _) => t
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
   659
      | Const (@{const_name Finite_Set.finite}, _) => t
35092
cfe605c54e50 moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents: 34974
diff changeset
   660
      | Const (@{const_name Orderings.less}, Type ("fun", [Type ("nat", []),
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   661
        Type ("fun", [Type ("nat", []), Type ("bool", [])])])) => t
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35092
diff changeset
   662
      | Const (@{const_name Groups.plus}, Type ("fun", [Type ("nat", []),
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   663
        Type ("fun", [Type ("nat", []), Type ("nat", [])])])) => t
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35092
diff changeset
   664
      | Const (@{const_name Groups.minus}, Type ("fun", [Type ("nat", []),
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   665
        Type ("fun", [Type ("nat", []), Type ("nat", [])])])) => t
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35092
diff changeset
   666
      | Const (@{const_name Groups.times}, Type ("fun", [Type ("nat", []),
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   667
        Type ("fun", [Type ("nat", []), Type ("nat", [])])])) => t
29820
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
   668
      | Const (@{const_name List.append}, _) => t
36130
9a672f7d488d commented out unsound "lfp"/"gfp" handling + fixed set output syntax;
blanchet
parents: 35746
diff changeset
   669
(* UNSOUND
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   670
      | Const (@{const_name lfp}, _) => t
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   671
      | Const (@{const_name gfp}, _) => t
36130
9a672f7d488d commented out unsound "lfp"/"gfp" handling + fixed set output syntax;
blanchet
parents: 35746
diff changeset
   672
*)
29820
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
   673
      | Const (@{const_name fst}, _) => t
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
   674
      | Const (@{const_name snd}, _) => t
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   675
      (* simply-typed lambda calculus *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   676
      | Const (s, T) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   677
        (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
   678
          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
   679
          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
   680
        (* 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
   681
        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
   682
          SOME (axname, rhs) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   683
          (* 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
   684
          (* 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
   685
          (* '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
   686
          (* 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
   687
          (* 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
   688
          (* 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
   689
          (* which could interact with overloading to create loops.  *)
32950
5d5e123443b3 normalized aliases of Output operations;
wenzelm
parents: 32857
diff changeset
   690
          ((*tracing (" unfolding: " ^ axname);*)
29820
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
   691
           unfold_loop rhs)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   692
        | NONE => t)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   693
      | Free _           => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   694
      | Var _            => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   695
      | Bound _          => t
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   696
      | 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
   697
      | 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
   698
    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
   699
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   700
    result
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   701
  end;
21985
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   702
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   703
(* ------------------------------------------------------------------------- *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   704
(* collect_axioms: collects (monomorphic, universally quantified, unfolded   *)
acfb13e8819e constants are unfolded, universal quantifiers are stripped, some minor changes
webertj
parents: 21931
diff changeset
   705
(*                 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
   706
(* ------------------------------------------------------------------------- *)
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
   707
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   708
  (* 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
   709
  (*       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
   710
  (*       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
   711
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   712
  (* 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
   713
  (*       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
   714
  (*       "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
   715
  (*       *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
   716
  (*       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
   717
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   718
  (* 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
   719
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   720
  (* 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
   721
  (* 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
   722
  (* 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
   723
  (* 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
   724
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   725
  (* 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
   726
  (* 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
   727
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   728
  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
   729
  let
32950
5d5e123443b3 normalized aliases of Output operations;
wenzelm
parents: 32857
diff changeset
   730
    val _ = tracing "Adding axioms..."
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   731
    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
   732
    fun collect_this_axiom (axname, ax) axs =
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   733
      let
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   734
        val ax' = unfold_defs thy ax
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   735
      in
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   736
        if member (op aconv) axs ax' then axs
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   737
        else (tracing axname; collect_term_axioms ax' (ax' :: axs))
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   738
      end
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   739
    and collect_sort_axioms T axs =
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   740
      let
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   741
        val sort =
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   742
          (case T of
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   743
            TFree (_, sort) => sort
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   744
          | TVar (_, sort)  => sort
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   745
          | _ => raise REFUTE ("collect_axioms",
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   746
              "type " ^ Syntax.string_of_typ_global thy T ^ " is not a variable"))
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   747
        (* obtain axioms for all superclasses *)
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   748
        val superclasses = sort @ maps (Sign.super_classes thy) sort
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   749
        (* merely an optimization, because 'collect_this_axiom' disallows *)
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   750
        (* duplicate axioms anyway:                                       *)
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   751
        val superclasses = distinct (op =) superclasses
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   752
        val class_axioms = maps (fn class => map (fn ax =>
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   753
          ("<" ^ class ^ ">", Thm.prop_of ax))
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   754
          (#axioms (AxClass.get_info thy class) handle ERROR _ => []))
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   755
          superclasses
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   756
        (* replace the (at most one) schematic type variable in each axiom *)
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   757
        (* by the actual type 'T'                                          *)
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   758
        val monomorphic_class_axioms = map (fn (axname, ax) =>
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   759
          (case Term.add_tvars ax [] of
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   760
            [] => (axname, ax)
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   761
          | [(idx, S)] => (axname, monomorphic_term (Vartab.make [(idx, (S, T))]) ax)
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   762
          | _ =>
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   763
            raise REFUTE ("collect_axioms", "class axiom " ^ axname ^ " (" ^
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   764
              Syntax.string_of_term_global thy ax ^
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   765
              ") contains more than one type variable")))
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   766
          class_axioms
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   767
      in
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   768
        fold collect_this_axiom monomorphic_class_axioms axs
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   769
      end
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   770
    and collect_type_axioms T axs =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   771
      case T of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   772
      (* simple types *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   773
        Type ("prop", []) => axs
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   774
      | Type ("fun", [T1, T2]) => collect_type_axioms T2 (collect_type_axioms T1 axs)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   775
      (* axiomatic type classes *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   776
      | Type ("itself", [T1]) => collect_type_axioms T1 axs
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   777
      | Type (s, Ts) =>
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31737
diff changeset
   778
        (case Datatype.get_info thy s of
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   779
          SOME info =>  (* inductive datatype *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   780
            (* only collect relevant type axioms for the argument types *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   781
            fold collect_type_axioms Ts axs
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   782
        | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   783
          (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
   784
            SOME (axname, ax) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   785
            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
   786
          | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   787
            (* 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
   788
            (* at least collect relevant type axioms for the argument types *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   789
            fold collect_type_axioms Ts axs))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   790
      (* axiomatic type classes *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   791
      | TFree _ => collect_sort_axioms T axs
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   792
      (* axiomatic type classes *)
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   793
      | TVar _ => collect_sort_axioms T axs
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   794
    and collect_term_axioms t axs =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   795
      case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   796
      (* Pure *)
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   797
        Const (@{const_name all}, _) => axs
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   798
      | Const (@{const_name "=="}, _) => axs
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   799
      | Const (@{const_name "==>"}, _) => axs
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   800
      (* axiomatic type classes *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   801
      | Const (@{const_name TYPE}, T) => collect_type_axioms T axs
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   802
      (* HOL *)
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   803
      | Const (@{const_name Trueprop}, _) => axs
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   804
      | Const (@{const_name Not}, _) => axs
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   805
      (* redundant, since 'True' is also an IDT constructor *)
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   806
      | Const (@{const_name True}, _) => axs
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   807
      (* redundant, since 'False' is also an IDT constructor *)
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   808
      | Const (@{const_name False}, _) => axs
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   809
      | Const (@{const_name undefined}, T) => collect_type_axioms T axs
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   810
      | Const (@{const_name The}, T) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   811
        let
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   812
          val ax = specialize_type thy (@{const_name The}, T)
29288
253bcf2a5854 avoid polymorphic equality;
wenzelm
parents: 29272
diff changeset
   813
            (the (AList.lookup (op =) axioms "HOL.the_eq_trivial"))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   814
        in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   815
          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
   816
        end
29820
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
   817
      | Const (@{const_name Hilbert_Choice.Eps}, T) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   818
        let
29820
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
   819
          val ax = specialize_type thy (@{const_name Hilbert_Choice.Eps}, T)
29288
253bcf2a5854 avoid polymorphic equality;
wenzelm
parents: 29272
diff changeset
   820
            (the (AList.lookup (op =) axioms "Hilbert_Choice.someI"))
22567
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
          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
   823
        end
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   824
      | Const (@{const_name All}, T) => collect_type_axioms T axs
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   825
      | Const (@{const_name Ex}, T) => collect_type_axioms T axs
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   826
      | Const (@{const_name "op ="}, T) => collect_type_axioms T axs
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   827
      | Const (@{const_name "op &"}, _) => axs
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   828
      | Const (@{const_name "op |"}, _) => axs
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   829
      | Const (@{const_name "op -->"}, _) => axs
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   830
      (* sets *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   831
      | Const (@{const_name Collect}, T) => collect_type_axioms T axs
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   832
      | Const (@{const_name "op :"}, T) => collect_type_axioms T axs
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   833
      (* other optimizations *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   834
      | Const (@{const_name Finite_Set.card}, T) => collect_type_axioms T axs
29820
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
   835
      | Const (@{const_name Finite_Set.finite}, T) =>
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   836
        collect_type_axioms T axs
35092
cfe605c54e50 moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents: 34974
diff changeset
   837
      | Const (@{const_name Orderings.less}, T as Type ("fun", [Type ("nat", []),
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   838
        Type ("fun", [Type ("nat", []), Type ("bool", [])])])) =>
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   839
          collect_type_axioms T axs
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35092
diff changeset
   840
      | Const (@{const_name Groups.plus}, T as Type ("fun", [Type ("nat", []),
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   841
        Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   842
          collect_type_axioms T axs
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35092
diff changeset
   843
      | Const (@{const_name Groups.minus}, T as Type ("fun", [Type ("nat", []),
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   844
        Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   845
          collect_type_axioms T axs
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35092
diff changeset
   846
      | Const (@{const_name Groups.times}, T as Type ("fun", [Type ("nat", []),
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   847
        Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   848
          collect_type_axioms T axs
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   849
      | Const (@{const_name List.append}, T) => collect_type_axioms T axs
36130
9a672f7d488d commented out unsound "lfp"/"gfp" handling + fixed set output syntax;
blanchet
parents: 35746
diff changeset
   850
(* UNSOUND
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   851
      | Const (@{const_name lfp}, T) => collect_type_axioms T axs
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   852
      | Const (@{const_name gfp}, T) => collect_type_axioms T axs
36130
9a672f7d488d commented out unsound "lfp"/"gfp" handling + fixed set output syntax;
blanchet
parents: 35746
diff changeset
   853
*)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   854
      | Const (@{const_name fst}, T) => collect_type_axioms T axs
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   855
      | Const (@{const_name snd}, T) => collect_type_axioms T axs
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   856
      (* simply-typed lambda calculus *)
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
   857
      | Const (s, T) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   858
          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
   859
            (* 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
   860
            (* and the class definition                               *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   861
            let
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   862
              val class = Logic.class_of_const s
31943
5e960a0780a2 renamed inclass/Inclass to of_class/OfClass, in accordance to of_sort;
wenzelm
parents: 31784
diff changeset
   863
              val of_class = Logic.mk_of_class (TVar (("'a", 0), [class]), class)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   864
              val ax_in = SOME (specialize_type thy (s, T) of_class)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   865
                (* 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
   866
                handle Type.TYPE_MATCH => NONE
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   867
              val ax_1 = Option.map (fn ax => (Syntax.string_of_term_global thy ax, ax)) ax_in
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   868
              val ax_2 = Option.map (apsnd (specialize_type thy (s, T))) (get_classdef thy class)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   869
            in
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   870
              collect_type_axioms T (fold collect_this_axiom (map_filter I [ax_1, ax_2]) axs)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   871
            end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   872
          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
   873
            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
   874
            (* only collect relevant type axioms *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   875
            collect_type_axioms T axs
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   876
          else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   877
            (* 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
   878
            (* 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
   879
            (* 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
   880
            (* only collect relevant type axioms *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   881
            collect_type_axioms T axs
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   882
      | Free (_, T) => collect_type_axioms T axs
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   883
      | Var (_, T) => collect_type_axioms T axs
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   884
      | Bound _ => axs
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   885
      | Abs (_, T, body) => collect_term_axioms body (collect_type_axioms T axs)
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   886
      | t1 $ t2 => collect_term_axioms t2 (collect_term_axioms t1 axs)
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   887
    val result = map close_form (collect_term_axioms t [])
32950
5d5e123443b3 normalized aliases of Output operations;
wenzelm
parents: 32857
diff changeset
   888
    val _ = tracing " ...done."
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   889
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   890
    result
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   891
  end;
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   892
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
   893
(* ------------------------------------------------------------------------- *)
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
   894
(* 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
   895
(*               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
   896
(*               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
   897
(*               'propT'.  For IDTs, also the argument types of constructors *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   898
(*               and all mutually recursive IDTs are considered.             *)
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
   899
(* ------------------------------------------------------------------------- *)
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
   900
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   901
  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
   902
  let
29272
fb3ccf499df5 use regular Term.add_XXX etc.;
wenzelm
parents: 29265
diff changeset
   903
    fun collect_types T acc =
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   904
      (case T of
29272
fb3ccf499df5 use regular Term.add_XXX etc.;
wenzelm
parents: 29265
diff changeset
   905
        Type ("fun", [T1, T2]) => collect_types T1 (collect_types T2 acc)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   906
      | Type ("prop", [])      => acc
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   907
      | Type (s, Ts)           =>
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31737
diff changeset
   908
        (case Datatype.get_info thy s of
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   909
          SOME info =>  (* inductive datatype *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   910
          let
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   911
            val index        = #index info
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   912
            val descr        = #descr info
29288
253bcf2a5854 avoid polymorphic equality;
wenzelm
parents: 29272
diff changeset
   913
            val (_, typs, _) = the (AList.lookup (op =) descr index)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   914
            val typ_assoc    = typs ~~ Ts
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   915
            (* sanity check: every element in 'dtyps' must be a *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   916
            (* 'DtTFree'                                        *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   917
            val _ = if Library.exists (fn d =>
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
   918
              case d of Datatype_Aux.DtTFree _ => false | _ => true) typs then
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   919
              raise REFUTE ("ground_types", "datatype argument (for type "
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26931
diff changeset
   920
                ^ Syntax.string_of_typ_global thy T ^ ") is not a variable")
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   921
            else ()
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   922
            (* required for mutually recursive datatypes; those need to   *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   923
            (* be added even if they are an instance of an otherwise non- *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   924
            (* recursive datatype                                         *)
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   925
            fun collect_dtyp d acc =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   926
            let
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   927
              val dT = typ_of_dtyp descr typ_assoc d
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   928
            in
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   929
              case d of
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
   930
                Datatype_Aux.DtTFree _ =>
29272
fb3ccf499df5 use regular Term.add_XXX etc.;
wenzelm
parents: 29265
diff changeset
   931
                collect_types dT acc
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
   932
              | Datatype_Aux.DtType (_, ds) =>
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   933
                collect_types dT (fold_rev collect_dtyp ds acc)
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
   934
              | Datatype_Aux.DtRec i =>
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   935
                if dT mem acc then
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   936
                  acc  (* prevent infinite recursion *)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   937
                else
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   938
                  let
29288
253bcf2a5854 avoid polymorphic equality;
wenzelm
parents: 29272
diff changeset
   939
                    val (_, dtyps, dconstrs) = the (AList.lookup (op =) descr i)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   940
                    (* if the current type is a recursive IDT (i.e. a depth *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   941
                    (* is required), add it to 'acc'                        *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   942
                    val acc_dT = if Library.exists (fn (_, ds) =>
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
   943
                      Library.exists Datatype_Aux.is_rec_type ds) dconstrs then
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   944
                        insert (op =) dT acc
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   945
                      else acc
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   946
                    (* collect argument types *)
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   947
                    val acc_dtyps = fold_rev collect_dtyp dtyps acc_dT
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   948
                    (* collect constructor types *)
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   949
                    val acc_dconstrs = fold_rev collect_dtyp (maps snd dconstrs) acc_dtyps
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   950
                  in
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   951
                    acc_dconstrs
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   952
                  end
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   953
            end
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   954
          in
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   955
            (* argument types 'Ts' could be added here, but they are also *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   956
            (* added by 'collect_dtyp' automatically                      *)
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
   957
            collect_dtyp (Datatype_Aux.DtRec index) acc
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   958
          end
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   959
        | NONE =>
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   960
          (* not an inductive datatype, e.g. defined via "typedef" or *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   961
          (* "typedecl"                                               *)
29272
fb3ccf499df5 use regular Term.add_XXX etc.;
wenzelm
parents: 29265
diff changeset
   962
          insert (op =) T (fold collect_types Ts acc))
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   963
      | TFree _                => insert (op =) T acc
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
   964
      | TVar _                 => insert (op =) T acc)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   965
  in
29272
fb3ccf499df5 use regular Term.add_XXX etc.;
wenzelm
parents: 29265
diff changeset
   966
    fold_types collect_types t []
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   967
  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
   968
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
   969
(* ------------------------------------------------------------------------- *)
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
   970
(* 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
   971
(*                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
   972
(*                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
   973
(*                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
   974
(* ------------------------------------------------------------------------- *)
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
   975
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   976
  (* 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
   977
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   978
  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
   979
    | 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
   980
    | 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
   981
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
   982
(* ------------------------------------------------------------------------- *)
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
   983
(* 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
   984
(*                 '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
   985
(*                 '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
   986
(* ------------------------------------------------------------------------- *)
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
   987
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   988
  (* 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
   989
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   990
  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
   991
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   992
    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
   993
      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
   994
        SOME n => n
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
   995
      | NONE => minsize
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   996
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
   997
    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
   998
  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
   999
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
  1000
(* ------------------------------------------------------------------------- *)
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
  1001
(* 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
  1002
(*                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
  1003
(*                '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
  1004
(*                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
  1005
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1006
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1007
  (* (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
  1008
    (Term.typ * int) list option *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1009
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1010
  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
  1011
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1012
    (* 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
  1013
    (* 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
  1014
    (* 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
  1015
    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
  1016
      if sum=0 then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1017
        SOME []
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1018
      else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1019
        NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1020
      | 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
  1021
      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
  1022
        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
  1023
      else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1024
        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
  1025
    (* 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
  1026
    (* 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
  1027
    (* 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
  1028
    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
  1029
      NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1030
      | 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
  1031
      (* 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
  1032
      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
  1033
      | 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
  1034
      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
  1035
        (* we can shift *)
33035
15eab423e573 standardized basic operations on type option;
wenzelm
parents: 33002
diff changeset
  1036
        SOME (the (make_first max (len+1) (sum+x1-1)) @ (x2+1) :: xs)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1037
      else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1038
        (* continue search *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1039
        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
  1040
    (* only consider those types for which the size is not fixed *)
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 33246
diff changeset
  1041
    val mutables = filter_out (AList.defined (op =) sizes o string_of_typ o fst) xs
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1042
    (* 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
  1043
    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
  1044
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1045
    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
  1046
      SOME diffs' =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1047
      (* merge with those types for which the size is fixed *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  1048
      SOME (fst (fold_map (fn (T, _) => fn ds =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1049
        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
  1050
        (* return the fixed size *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  1051
          SOME n => ((T, n), ds)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1052
        (* consume the head of 'ds', add 'minsize' *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  1053
        | NONE   => ((T, minsize + hd ds), tl ds))
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  1054
        xs diffs'))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1055
    | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1056
      NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1057
  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
  1058
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
  1059
(* ------------------------------------------------------------------------- *)
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
  1060
(* 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
  1061
(*         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
  1062
(* ------------------------------------------------------------------------- *)
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
  1063
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1064
  (* 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
  1065
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1066
  fun toTrue (Leaf [fm, _]) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1067
    fm
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1068
    | toTrue _              =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1069
    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
  1070
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
  1071
(* ------------------------------------------------------------------------- *)
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
  1072
(* 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
  1073
(*          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
  1074
(*          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
  1075
(* ------------------------------------------------------------------------- *)
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
  1076
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1077
  (* 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
  1078
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1079
  fun toFalse (Leaf [_, fm]) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1080
    fm
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1081
    | toFalse _              =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1082
    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
  1083
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
  1084
(* ------------------------------------------------------------------------- *)
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
  1085
(* 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
  1086
(*             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
  1087
(*             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
  1088
(* 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
  1089
(* {...}     : parameters that control the translation/model generation      *)
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1090
(* assm_ts   : assumptions to be considered unless "no_assms" is specified   *)
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
  1091
(* 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
  1092
(* 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
  1093
(* ------------------------------------------------------------------------- *)
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
  1094
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1095
  (* 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
  1096
30314
853778f6ef7d Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
blanchet
parents: 30312
diff changeset
  1097
  fun find_model thy {sizes, minsize, maxsize, maxvars, maxtime, satsolver,
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1098
    no_assms, expect} assm_ts t negate =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1099
  let
33054
dd1192a96968 fixed the "expect" mechanism of Refute in the face of timeouts
blanchet
parents: 32952
diff changeset
  1100
    (* string -> unit *)
dd1192a96968 fixed the "expect" mechanism of Refute in the face of timeouts
blanchet
parents: 32952
diff changeset
  1101
    fun check_expect outcome_code =
dd1192a96968 fixed the "expect" mechanism of Refute in the face of timeouts
blanchet
parents: 32952
diff changeset
  1102
      if expect = "" orelse outcome_code = expect then ()
dd1192a96968 fixed the "expect" mechanism of Refute in the face of timeouts
blanchet
parents: 32952
diff changeset
  1103
      else error ("Unexpected outcome: " ^ quote outcome_code ^ ".")
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1104
    (* unit -> unit *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1105
    fun wrapper () =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1106
    let
30349
110826d1e5ff Added a second timeout mechanism to Refute.
blanchet
parents: 30347
diff changeset
  1107
      val timer  = Timer.startRealTimer ()
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1108
      val t = if no_assms then t
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1109
              else if negate then Logic.list_implies (assm_ts, t)
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1110
              else Logic.mk_conjunction_list (t :: assm_ts)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1111
      val u      = unfold_defs thy t
32950
5d5e123443b3 normalized aliases of Output operations;
wenzelm
parents: 32857
diff changeset
  1112
      val _      = tracing ("Unfolded term: " ^ Syntax.string_of_term_global thy u)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1113
      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
  1114
      (* Term.typ list *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  1115
      val types = fold (union (op =) o ground_types thy) (u :: axioms) []
32950
5d5e123443b3 normalized aliases of Output operations;
wenzelm
parents: 32857
diff changeset
  1116
      val _     = tracing ("Ground types: "
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1117
        ^ (if null types then "none."
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26931
diff changeset
  1118
           else commas (map (Syntax.string_of_typ_global thy) types)))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1119
      (* 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
  1120
      (* 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
  1121
      (* TODO: no warning needed for /positive/ occurrences of IDTs       *)
30314
853778f6ef7d Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
blanchet
parents: 30312
diff changeset
  1122
      val maybe_spurious = Library.exists (fn
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1123
          Type (s, _) =>
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31737
diff changeset
  1124
          (case Datatype.get_info thy s of
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1125
            SOME info =>  (* inductive datatype *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1126
            let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1127
              val index           = #index info
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1128
              val descr           = #descr info
29288
253bcf2a5854 avoid polymorphic equality;
wenzelm
parents: 29272
diff changeset
  1129
              val (_, _, constrs) = the (AList.lookup (op =) descr index)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1130
            in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1131
              (* recursive datatype? *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1132
              Library.exists (fn (_, ds) =>
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
  1133
                Library.exists Datatype_Aux.is_rec_type ds) constrs
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1134
            end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1135
          | NONE => false)
30314
853778f6ef7d Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
blanchet
parents: 30312
diff changeset
  1136
        | _ => false) types
853778f6ef7d Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
blanchet
parents: 30312
diff changeset
  1137
      val _ = if maybe_spurious then
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1138
          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
  1139
            ^ "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
  1140
        else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1141
          ()
30314
853778f6ef7d Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
blanchet
parents: 30312
diff changeset
  1142
      (* (Term.typ * int) list -> string *)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1143
      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
  1144
      let
30349
110826d1e5ff Added a second timeout mechanism to Refute.
blanchet
parents: 30347
diff changeset
  1145
        val msecs_spent = Time.toMilliseconds (Timer.checkRealTimer timer)
110826d1e5ff Added a second timeout mechanism to Refute.
blanchet
parents: 30347
diff changeset
  1146
        val _ = maxtime = 0 orelse msecs_spent < 1000 * maxtime
110826d1e5ff Added a second timeout mechanism to Refute.
blanchet
parents: 30347
diff changeset
  1147
                orelse raise TimeLimit.TimeOut
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1148
        val init_model = (universe, [])
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1149
        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
  1150
          bounds = [], wellformed = True}
32950
5d5e123443b3 normalized aliases of Output operations;
wenzelm
parents: 32857
diff changeset
  1151
        val _ = tracing ("Translating term (sizes: "
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1152
          ^ 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
  1153
        (* translate 'u' and all axioms *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  1154
        val (intrs, (model, args)) = fold_map (fn t' => fn (m, a) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1155
          let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1156
            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
  1157
          in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1158
            (* set 'def_eq' to 'true' *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  1159
            (i, (m', {maxvars = #maxvars a', def_eq = true,
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1160
              next_idx = #next_idx a', bounds = #bounds a',
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  1161
              wellformed = #wellformed a'}))
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  1162
          end) (u :: axioms) (init_model, init_args)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1163
        (* 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
  1164
        (* 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
  1165
        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
  1166
        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
  1167
        val fm    = PropLogic.all [#wellformed args, fm_ax, fm_u]
34017
ef2776c89799 better error message in Refute when specifying a non-existing SAT solver
blanchet
parents: 33968
diff changeset
  1168
        val solver =
ef2776c89799 better error message in Refute when specifying a non-existing SAT solver
blanchet
parents: 33968
diff changeset
  1169
          SatSolver.invoke_solver satsolver
ef2776c89799 better error message in Refute when specifying a non-existing SAT solver
blanchet
parents: 33968
diff changeset
  1170
          handle Option.Option =>
ef2776c89799 better error message in Refute when specifying a non-existing SAT solver
blanchet
parents: 33968
diff changeset
  1171
                 error ("Unknown SAT solver: " ^ quote satsolver ^
ef2776c89799 better error message in Refute when specifying a non-existing SAT solver
blanchet
parents: 33968
diff changeset
  1172
                        ". Available solvers: " ^
ef2776c89799 better error message in Refute when specifying a non-existing SAT solver
blanchet
parents: 33968
diff changeset
  1173
                        commas (map (quote o fst) (!SatSolver.solvers)) ^ ".")
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1174
      in
32950
5d5e123443b3 normalized aliases of Output operations;
wenzelm
parents: 32857
diff changeset
  1175
        priority "Invoking SAT solver...";
34017
ef2776c89799 better error message in Refute when specifying a non-existing SAT solver
blanchet
parents: 33968
diff changeset
  1176
        (case solver fm of
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1177
          SatSolver.SATISFIABLE assignment =>
32950
5d5e123443b3 normalized aliases of Output operations;
wenzelm
parents: 32857
diff changeset
  1178
          (priority ("*** Model found: ***\n" ^ print_model thy model
30314
853778f6ef7d Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
blanchet
parents: 30312
diff changeset
  1179
            (fn i => case assignment i of SOME b => b | NONE => true));
30347
91f73b2997f9 Refute: Distinguish between "genuine" and "potential" in the newly added "expect" option.
blanchet
parents: 30314
diff changeset
  1180
           if maybe_spurious then "potential" else "genuine")
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1181
        | SatSolver.UNSATISFIABLE _ =>
32950
5d5e123443b3 normalized aliases of Output operations;
wenzelm
parents: 32857
diff changeset
  1182
          (priority "No model exists.";
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1183
          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
  1184
            SOME universe' => find_model_loop universe'
32950
5d5e123443b3 normalized aliases of Output operations;
wenzelm
parents: 32857
diff changeset
  1185
          | NONE           => (priority
30314
853778f6ef7d Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
blanchet
parents: 30312
diff changeset
  1186
            "Search terminated, no larger universe within the given limits.";
853778f6ef7d Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
blanchet
parents: 30312
diff changeset
  1187
            "none"))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1188
        | SatSolver.UNKNOWN =>
32950
5d5e123443b3 normalized aliases of Output operations;
wenzelm
parents: 32857
diff changeset
  1189
          (priority "No model found.";
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1190
          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
  1191
            SOME universe' => find_model_loop universe'
32950
5d5e123443b3 normalized aliases of Output operations;
wenzelm
parents: 32857
diff changeset
  1192
          | NONE           => (priority
30314
853778f6ef7d Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
blanchet
parents: 30312
diff changeset
  1193
            "Search terminated, no larger universe within the given limits.";
853778f6ef7d Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
blanchet
parents: 30312
diff changeset
  1194
            "unknown"))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1195
        ) handle SatSolver.NOT_CONFIGURED =>
30314
853778f6ef7d Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
blanchet
parents: 30312
diff changeset
  1196
          (error ("SAT solver " ^ quote satsolver ^ " is not configured.");
853778f6ef7d Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
blanchet
parents: 30312
diff changeset
  1197
           "unknown")
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1198
      end handle MAXVARS_EXCEEDED =>
32950
5d5e123443b3 normalized aliases of Output operations;
wenzelm
parents: 32857
diff changeset
  1199
        (priority ("Search terminated, number of Boolean variables ("
30314
853778f6ef7d Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
blanchet
parents: 30312
diff changeset
  1200
          ^ string_of_int maxvars ^ " allowed) exceeded.");
853778f6ef7d Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
blanchet
parents: 30312
diff changeset
  1201
          "unknown")
853778f6ef7d Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
blanchet
parents: 30312
diff changeset
  1202
        val outcome_code = find_model_loop (first_universe types sizes minsize)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1203
      in
33054
dd1192a96968 fixed the "expect" mechanism of Refute in the face of timeouts
blanchet
parents: 32952
diff changeset
  1204
        check_expect outcome_code
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1205
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1206
    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1207
      (* some parameter sanity checks *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1208
      minsize>=1 orelse
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1209
        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
  1210
      maxsize>=1 orelse
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1211
        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
  1212
      maxsize>=minsize orelse
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1213
        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
  1214
        ") 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
  1215
      maxvars>=0 orelse
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1216
        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
  1217
      maxtime>=0 orelse
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1218
        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
  1219
      (* enter loop with or without time limit *)
32950
5d5e123443b3 normalized aliases of Output operations;
wenzelm
parents: 32857
diff changeset
  1220
      priority ("Trying to find a model that "
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1221
        ^ (if negate then "refutes" else "satisfies") ^ ": "
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26931
diff changeset
  1222
        ^ Syntax.string_of_term_global thy t);
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1223
      if maxtime>0 then (
24688
a5754ca5c510 replaced interrupt_timeout by TimeLimit.timeLimit (available on SML/NJ and Poly/ML 5.1);
wenzelm
parents: 23881
diff changeset
  1224
        TimeLimit.timeLimit (Time.fromSeconds maxtime)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1225
          wrapper ()
24688
a5754ca5c510 replaced interrupt_timeout by TimeLimit.timeLimit (available on SML/NJ and Poly/ML 5.1);
wenzelm
parents: 23881
diff changeset
  1226
        handle TimeLimit.TimeOut =>
33054
dd1192a96968 fixed the "expect" mechanism of Refute in the face of timeouts
blanchet
parents: 32952
diff changeset
  1227
          (priority ("Search terminated, time limit (" ^
dd1192a96968 fixed the "expect" mechanism of Refute in the face of timeouts
blanchet
parents: 32952
diff changeset
  1228
              string_of_int maxtime
dd1192a96968 fixed the "expect" mechanism of Refute in the face of timeouts
blanchet
parents: 32952
diff changeset
  1229
              ^ (if maxtime=1 then " second" else " seconds") ^ ") exceeded.");
dd1192a96968 fixed the "expect" mechanism of Refute in the face of timeouts
blanchet
parents: 32952
diff changeset
  1230
           check_expect "unknown")
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1231
      ) else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1232
        wrapper ()
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1233
    end;
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1234
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1235
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1236
(* ------------------------------------------------------------------------- *)
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1237
(* INTERFACE, PART 2: FINDING A MODEL                                        *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1238
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1239
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1240
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1241
(* 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
  1242
(* 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
  1243
(*               parameters                                                  *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1244
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1245
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1246
  (* theory -> (string * string) list -> Term.term list -> Term.term -> unit *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1247
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1248
  fun satisfy_term thy params assm_ts t =
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1249
    find_model thy (actual_params thy params) assm_ts t false;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1250
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1251
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  1252
(* 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
  1253
(* 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
  1254
(*              parameters                                                   *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1255
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1256
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1257
  (* theory -> (string * string) list -> Term.term list -> Term.term -> unit *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1258
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1259
  fun refute_term thy params assm_ts t =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1260
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1261
    (* 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
  1262
    (* 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
  1263
    (* 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
  1264
    (* for ALL types, not ...)                                             *)
29272
fb3ccf499df5 use regular Term.add_XXX etc.;
wenzelm
parents: 29265
diff changeset
  1265
    val _ = null (Term.add_tvars t []) orelse
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1266
      error "Term to be refuted contains schematic type variables"
21556
e0ffb2d13f9f outermost universal quantifiers are stripped
webertj
parents: 21267
diff changeset
  1267
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1268
    (* 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
  1269
    (* (Term.indexname * Term.typ) list *)
29265
5b4247055bd7 moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents: 29004
diff changeset
  1270
    val vars = sort_wrt (fst o fst) (map dest_Var (OldTerm.term_vars t))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1271
    (* Term.term *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  1272
    val ex_closure = fold (fn ((x, i), T) => fn t' =>
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  1273
      HOLogic.exists_const T $
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  1274
        Abs (x, T, abstract_over (Var ((x, i), T), t'))) vars t
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1275
    (* 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
  1276
    (* '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
  1277
    (* 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
  1278
    (* resulting term correctly, without checking its type.              *)
21556
e0ffb2d13f9f outermost universal quantifiers are stripped
webertj
parents: 21267
diff changeset
  1279
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1280
    (* 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
  1281
    (* 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
  1282
    (* 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
  1283
    (* 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
  1284
    (* 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
  1285
    (* 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
  1286
    (* quantified variables.                                             *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1287
    (* maps  !!x1...xn. !xk...xm. t   to   t  *)
29820
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
  1288
    fun strip_all_body (Const (@{const_name all}, _) $ Abs (_, _, t)) =
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
  1289
        strip_all_body t
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
  1290
      | strip_all_body (Const (@{const_name Trueprop}, _) $ t) =
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
  1291
        strip_all_body t
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
  1292
      | strip_all_body (Const (@{const_name All}, _) $ Abs (_, _, t)) =
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
  1293
        strip_all_body t
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1294
      | strip_all_body t = t
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1295
    (* maps  !!x1...xn. !xk...xm. t   to   [x1, ..., xn, xk, ..., xm]  *)
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1296
    fun strip_all_vars (Const (@{const_name all}, _) $ Abs (a, T, t)) =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1297
      (a, T) :: strip_all_vars t
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1298
      | strip_all_vars (Const (@{const_name Trueprop}, _) $ t) =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1299
      strip_all_vars t
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1300
      | strip_all_vars (Const (@{const_name All}, _) $ Abs (a, T, t)) =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1301
      (a, T) :: strip_all_vars t
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1302
      | strip_all_vars t =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1303
      [] : (string * typ) list
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1304
    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
  1305
    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
  1306
    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
  1307
  in
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1308
    find_model thy (actual_params thy params) assm_ts subst_t true
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1309
  end;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1310
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1311
(* ------------------------------------------------------------------------- *)
32857
394d37f19e0a Refute.refute_goal: canonical goal addresses from 1 (renamed from refute_subgoal to clarify change in semantics);
wenzelm
parents: 31986
diff changeset
  1312
(* refute_goal                                                               *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1313
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1314
34120
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1315
  fun refute_goal ctxt params th i =
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1316
  let
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1317
    val t = th |> prop_of
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1318
  in
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1319
    if Logic.count_prems t = 0 then
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1320
      priority "No subgoal!"
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1321
    else
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1322
      let
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1323
        val assms = map term_of (Assumption.all_assms_of ctxt)
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1324
        val (t, frees) = Logic.goal_params t i
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1325
      in
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1326
        refute_term (ProofContext.theory_of ctxt) params assms
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1327
        (subst_bounds (frees, t))
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1328
      end
f9920a3ddf50 added "no_assms" option to Refute, and include structured proof assumptions by default;
blanchet
parents: 34017
diff changeset
  1329
  end
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1330
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1331
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1332
(* ------------------------------------------------------------------------- *)
15292
09e218879265 minor changes (comments/code refactoring)
webertj
parents: 15283
diff changeset
  1333
(* INTERPRETERS: Auxiliary Functions                                         *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1334
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1335
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1336
(* ------------------------------------------------------------------------- *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1337
(* make_constants: returns all interpretations for type 'T' that consist of  *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1338
(*                 unit vectors with 'True'/'False' only (no Boolean         *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1339
(*                 variables)                                                *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1340
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1341
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1342
  (* theory -> model -> Term.typ -> interpretation list *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  1343
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1344
  fun make_constants thy model T =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1345
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1346
    (* 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
  1347
    (* int -> interpretation list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1348
    fun unit_vectors n =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1349
    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1350
      (* 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
  1351
      (* int * int -> interpretation *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1352
      fun unit_vector (k, n) =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1353
        Leaf ((replicate (k-1) False) @ (True :: (replicate (n-k) False)))
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1354
      (* int -> interpretation list *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1355
      fun unit_vectors_loop k =
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1356
        if k>n then [] else unit_vector (k,n) :: unit_vectors_loop (k+1)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1357
    in
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1358
      unit_vectors_loop 1
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1359
    end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1360
    (* 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
  1361
    (* identical) elements from 'xs'                               *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1362
    (* 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
  1363
    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
  1364
      map single xs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1365
      | pick_all n xs =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1366
      let val rec_pick = pick_all (n-1) xs in
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32950
diff changeset
  1367
        maps (fn x => map (cons x) rec_pick) xs
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1368
      end
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1369
    (* returns all constant interpretations that have the same tree *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1370
    (* structure as the interpretation argument                     *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1371
    (* interpretation -> interpretation list *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1372
    fun make_constants_intr (Leaf xs) = unit_vectors (length xs)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1373
      | make_constants_intr (Node xs) = map Node (pick_all (length xs)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1374
      (make_constants_intr (hd xs)))
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1375
    (* obtain the interpretation for a variable of type 'T' *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1376
    val (i, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1,
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1377
      bounds=[], wellformed=True} (Free ("dummy", T))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1378
  in
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1379
    make_constants_intr i
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1380
  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
  1381
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
  1382
(* ------------------------------------------------------------------------- *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1383
(* power: 'power (a, b)' computes a^b, for a>=0, b>=0                        *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1384
(* ------------------------------------------------------------------------- *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1385
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1386
  (* int * int -> int *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1387
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1388
  fun power (a, 0) = 1
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1389
    | power (a, 1) = a
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1390
    | power (a, b) = let val ab = power(a, b div 2) in
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1391
        ab * ab * power(a, b mod 2)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1392
      end;
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1393
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1394
(* ------------------------------------------------------------------------- *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1395
(* size_of_type: returns the number of elements in a type 'T' (i.e. 'length  *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1396
(*               (make_constants T)', but implemented more efficiently)      *)
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
  1397
(* ------------------------------------------------------------------------- *)
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
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1399
  (* theory -> model -> Term.typ -> 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
  1400
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1401
  (* returns 0 for an empty ground type or a function type with empty      *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1402
  (* codomain, but fails for a function type with empty domain --          *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1403
  (* admissibility of datatype constructor argument types (see "Inductive  *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1404
  (* datatypes in HOL - lessons learned ...", S. Berghofer, M. Wenzel,     *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1405
  (* TPHOLs 99) ensures that recursive, possibly empty, datatype fragments *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1406
  (* never occur as the domain of a function type that is the type of a    *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1407
  (* constructor argument                                                  *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1408
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1409
  fun size_of_type thy model T =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1410
  let
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1411
    (* returns the number of elements that have the same tree structure as a *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1412
    (* given interpretation                                                  *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1413
    fun size_of_intr (Leaf xs) = length xs
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1414
      | size_of_intr (Node xs) = power (size_of_intr (hd xs), length xs)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1415
    (* obtain the interpretation for a variable of type 'T' *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1416
    val (i, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1,
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1417
      bounds=[], wellformed=True} (Free ("dummy", T))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1418
  in
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1419
    size_of_intr i
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1420
  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
  1421
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
  1422
(* ------------------------------------------------------------------------- *)
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
(* 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
  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
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1426
  (* 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
  1427
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1428
  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
  1429
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1430
  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
  1431
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
  1432
(* ------------------------------------------------------------------------- *)
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
(* 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
  1434
(*                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
  1435
(* - 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
  1436
(*   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
  1437
(* - 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
  1438
(*   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
  1439
(* - 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
  1440
(*   '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
  1441
(* ------------------------------------------------------------------------- *)
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
  1442
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1443
  (* 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
  1444
  (* 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
  1445
  (* 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
  1446
  (* '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
  1447
  (* 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
  1448
  (* 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
  1449
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1450
  (* 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
  1451
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1452
  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
  1453
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1454
    (* interpretation * interpretation -> prop_formula *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1455
    fun equal (i1, i2) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1456
      (case i1 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1457
        Leaf xs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1458
        (case i2 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1459
          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
  1460
        | 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
  1461
          "second interpretation is higher"))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1462
      | Node 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
          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
  1465
          "first interpretation is higher")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1466
        | 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
  1467
    (* interpretation * interpretation -> prop_formula *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1468
    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
  1469
      (case i1 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1470
        Leaf 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
          (* defined and not equal *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1473
          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
  1474
          :: (PropLogic.exists ys)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1475
          :: (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
  1476
        | 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
  1477
          "second interpretation is higher"))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1478
      | Node xs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1479
        (case i2 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1480
          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
  1481
          "first interpretation is higher")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1482
        | 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
  1483
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1484
    (* 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
  1485
    (* negation of 'equal'                                             *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1486
    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
  1487
  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
  1488
15292
09e218879265 minor changes (comments/code refactoring)
webertj
parents: 15283
diff changeset
  1489
(* ------------------------------------------------------------------------- *)
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
  1490
(* 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
  1491
(*                    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
  1492
(* 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
  1493
(* 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
  1494
(* 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
  1495
(* 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
  1496
(* ------------------------------------------------------------------------- *)
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
  1497
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1498
  (* 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
  1499
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1500
  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
  1501
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1502
    (* interpretation * interpretation -> prop_formula *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1503
    fun equal (i1, i2) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1504
      (case i1 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1505
        Leaf xs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1506
        (case i2 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1507
          (* 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
  1508
          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
  1509
          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
  1510
        | 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
  1511
          "second interpretation is higher"))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1512
      | Node xs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1513
        (case i2 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1514
          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
  1515
          "first interpretation is higher")
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1516
        | 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
  1517
    (* interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1518
    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
  1519
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1520
    Leaf [eq, SNot eq]
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1521
  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
  1522
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
  1523
(* ------------------------------------------------------------------------- *)
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
  1524
(* interpretation_apply: returns an interpretation that denotes the result   *)
22092
ab3dfcef6489 reformatted to 80 chars/line
webertj
parents: 22055
diff changeset
  1525
(*                       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
  1526
(*                       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
  1527
(* ------------------------------------------------------------------------- *)
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
  1528
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1529
  (* 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
  1530
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1531
  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
  1532
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1533
    (* interpretation * interpretation -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1534
    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
  1535
      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
  1536
        (tree_pair (tr1,tr2))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1537
    (* prop_formula * interpretation -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1538
    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
  1539
      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
  1540
    (* 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
  1541
    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
  1542
      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
  1543
      | 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
  1544
      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
  1545
        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
  1546
      | 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
  1547
      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
  1548
    (* 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
  1549
    (* lists                                                              *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1550
    (* '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
  1551
    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
  1552
      map (cons x) xss
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1553
    (* 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
  1554
    (* element of 'xss'                                                      *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1555
    (* '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
  1556
    fun pick_all [xs] =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1557
      map single xs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1558
      | pick_all (xs::xss) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1559
      let val rec_pick = pick_all xss in
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32950
diff changeset
  1560
        maps (fn x => map (cons x) rec_pick) xs
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1561
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1562
      | pick_all _ =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1563
      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
  1564
    (* interpretation -> prop_formula list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1565
    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
  1566
      xs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1567
      | 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
  1568
      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
  1569
        (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
  1570
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1571
    case i1 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1572
      Leaf _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1573
      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
  1574
    | Node xs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1575
      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
  1576
        (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
  1577
  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
  1578
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
  1579
(* ------------------------------------------------------------------------- *)
15292
09e218879265 minor changes (comments/code refactoring)
webertj
parents: 15283
diff changeset
  1580
(* eta_expand: eta-expands a term 't' by adding 'i' lambda abstractions      *)
09e218879265 minor changes (comments/code refactoring)
webertj
parents: 15283
diff changeset
  1581
(* ------------------------------------------------------------------------- *)
09e218879265 minor changes (comments/code refactoring)
webertj
parents: 15283
diff changeset
  1582
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1583
  (* Term.term -> int -> Term.term *)
15292
09e218879265 minor changes (comments/code refactoring)
webertj
parents: 15283
diff changeset
  1584
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1585
  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
  1586
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1587
    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
  1588
    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
  1589
  in
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
  1590
    fold_rev (fn T => fn term => Abs ("<eta_expand>", T, term))
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
  1591
      (List.take (Ts, i))
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
  1592
      (Term.list_comb (t', map Bound (i-1 downto 0)))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1593
  end;
15292
09e218879265 minor changes (comments/code refactoring)
webertj
parents: 15283
diff changeset
  1594
15335
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
  1595
(* ------------------------------------------------------------------------- *)
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
  1596
(* 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
  1597
(*               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
  1598
(*               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
  1599
(* ------------------------------------------------------------------------- *)
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
  1600
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1601
  fun size_of_dtyp thy typ_sizes descr typ_assoc constructors =
33002
f3f02f36a3e2 uniform use of Integer.add/mult/sum/prod;
wenzelm
parents: 32952
diff changeset
  1602
    Integer.sum (map (fn (_, dtyps) =>
f3f02f36a3e2 uniform use of Integer.add/mult/sum/prod;
wenzelm
parents: 32952
diff changeset
  1603
      Integer.prod (map (size_of_type thy (typ_sizes, []) o
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1604
        (typ_of_dtyp descr typ_assoc)) dtyps))
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1605
          constructors);
15335
f81e6e24351f minor code refactoring (typ_of_dtyp, size_of_dtyp)
webertj
parents: 15334
diff changeset
  1606
15292
09e218879265 minor changes (comments/code refactoring)
webertj
parents: 15283
diff changeset
  1607
09e218879265 minor changes (comments/code refactoring)
webertj
parents: 15283
diff changeset
  1608
(* ------------------------------------------------------------------------- *)
09e218879265 minor changes (comments/code refactoring)
webertj
parents: 15283
diff changeset
  1609
(* INTERPRETERS: Actual Interpreters                                         *)
09e218879265 minor changes (comments/code refactoring)
webertj
parents: 15283
diff changeset
  1610
(* ------------------------------------------------------------------------- *)
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
  1611
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1612
  (* 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
  1613
    (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
  1614
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1615
  (* 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
  1616
  (* 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
  1617
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1618
  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
  1619
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1620
    val (typs, terms)                                   = model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1621
    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
  1622
    (* 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
  1623
    fun interpret_groundterm T =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1624
    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1625
      (* 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
  1626
      fun interpret_groundtype () =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1627
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1628
        (* the model must specify a size for ground types *)
29820
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
  1629
        val size = if T = Term.propT then 2
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
  1630
          else the (AList.lookup (op =) typs T)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1631
        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
  1632
        (* 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
  1633
        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
  1634
          raise MAXVARS_EXCEEDED else ())
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1635
        (* prop_formula list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1636
        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
  1637
        (* interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1638
        val intr = Leaf fms
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1639
        (* 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
  1640
        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
  1641
          | 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
  1642
          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
  1643
        (* prop_formula *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1644
        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
  1645
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1646
        (* 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
  1647
        (* condition                                                  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1648
        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
  1649
          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
  1650
          wellformed = SAnd (wellformed, wf)})
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1651
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1652
    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1653
      case T of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1654
        Type ("fun", [T1, T2]) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1655
        let
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1656
          (* we create 'size_of_type ... T1' different copies of the        *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1657
          (* interpretation for 'T2', which are then combined into a single *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1658
          (* new interpretation                                             *)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1659
          (* 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
  1660
          (* 'idx': next variable index                         *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1661
          (* 'n'  : number of copies                            *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1662
          (* 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
  1663
          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
  1664
            (idx, [], True)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1665
            | make_copies idx n =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1666
            let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1667
              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
  1668
                {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
  1669
                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
  1670
              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
  1671
            in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1672
              (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
  1673
            end
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1674
          val (next, copies, wf) = make_copies next_idx
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1675
            (size_of_type thy model T1)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1676
          (* 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
  1677
          val intr = Node copies
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1678
        in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1679
          (* 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
  1680
          (* condition                                                  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1681
          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
  1682
            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
  1683
            wellformed = SAnd (wellformed, wf)})
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1684
        end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1685
      | Type _  => interpret_groundtype ()
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1686
      | TFree _ => interpret_groundtype ()
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1687
      | TVar  _ => interpret_groundtype ()
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1688
    end
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
    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
  1691
      SOME intr =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1692
      (* return an existing interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1693
      SOME (intr, model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1694
    | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1695
      (case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1696
        Const (_, T)     =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1697
        interpret_groundterm T
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1698
      | Free (_, T)      =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1699
        interpret_groundterm T
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1700
      | Var (_, T)       =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1701
        interpret_groundterm T
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1702
      | Bound i          =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1703
        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
  1704
      | Abs (x, T, body) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1705
        let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1706
          (* create all constants of type 'T' *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1707
          val constants = make_constants thy model T
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1708
          (* interpret the 'body' separately for each constant *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  1709
          val (bodies, (model', args')) = fold_map
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  1710
            (fn c => fn (m, a) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1711
              let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1712
                (* add 'c' to 'bounds' *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1713
                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
  1714
                  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
  1715
                  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
  1716
              in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1717
                (* 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
  1718
                (* but use old 'bounds'                                   *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  1719
                (i', (m', {maxvars = maxvars, def_eq = def_eq,
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1720
                  next_idx = #next_idx a', bounds = bounds,
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  1721
                  wellformed = #wellformed a'}))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1722
              end)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  1723
            constants (model, args)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1724
        in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1725
          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
  1726
        end
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  1727
      | t1 $ t2 =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1728
        let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1729
          (* 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
  1730
          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
  1731
          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
  1732
        in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1733
          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
  1734
        end)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1735
  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
  1736
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1737
  (* 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
  1738
    (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
  1739
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1740
  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
  1741
    case t of
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1742
      Const (@{const_name all}, _) $ t1 =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1743
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1744
        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
  1745
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1746
        case i of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1747
          Node xs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1748
          (* 3-valued logic *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1749
          let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1750
            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
  1751
            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
  1752
          in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1753
            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
  1754
          end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1755
        | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1756
          raise REFUTE ("Pure_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1757
            "\"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
  1758
      end
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1759
    | Const (@{const_name all}, _) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1760
      SOME (interpret thy model args (eta_expand t 1))
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1761
    | Const (@{const_name "=="}, _) $ t1 $ t2 =>
22567
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 (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
  1764
        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
  1765
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1766
        (* 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
  1767
        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
  1768
          (i1, i2), m2, a2)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1769
      end
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1770
    | Const (@{const_name "=="}, _) $ t1 =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1771
      SOME (interpret thy model args (eta_expand t 1))
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1772
    | Const (@{const_name "=="}, _) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1773
      SOME (interpret thy model args (eta_expand t 2))
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1774
    | Const (@{const_name "==>"}, _) $ t1 $ t2 =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1775
      (* 3-valued logic *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1776
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1777
        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
  1778
        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
  1779
        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
  1780
        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
  1781
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1782
        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
  1783
      end
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1784
    | Const (@{const_name "==>"}, _) $ t1 =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1785
      SOME (interpret thy model args (eta_expand t 1))
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1786
    | Const (@{const_name "==>"}, _) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1787
      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
  1788
    | _ => 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
  1789
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1790
  (* 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
  1791
    (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
  1792
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1793
  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
  1794
  (* 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
  1795
  (* 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
  1796
  (* 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
  1797
    case t of
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1798
      Const (@{const_name Trueprop}, _) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1799
      SOME (Node [TT, FF], model, args)
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1800
    | Const (@{const_name Not}, _) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1801
      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
  1802
    (* redundant, since 'True' is also an IDT constructor *)
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1803
    | Const (@{const_name True}, _) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1804
      SOME (TT, model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1805
    (* redundant, since 'False' is also an IDT constructor *)
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1806
    | Const (@{const_name False}, _) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1807
      SOME (FF, model, args)
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1808
    | Const (@{const_name All}, _) $ t1 =>  (* similar to "all" (Pure) *)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1809
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1810
        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
  1811
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1812
        case i of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1813
          Node xs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1814
          (* 3-valued logic *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1815
          let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1816
            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
  1817
            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
  1818
          in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1819
            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
  1820
          end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1821
        | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1822
          raise REFUTE ("HOLogic_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1823
            "\"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
  1824
      end
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1825
    | Const (@{const_name All}, _) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1826
      SOME (interpret thy model args (eta_expand t 1))
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1827
    | Const (@{const_name Ex}, _) $ t1 =>
22567
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.exists (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.all (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
            "\"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
  1843
      end
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1844
    | Const (@{const_name Ex}, _) =>
22567
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))
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1846
    | Const (@{const_name "op ="}, _) $ t1 $ t2 =>  (* similar to "==" (Pure) *)
22567
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 (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
  1849
        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
  1850
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1851
        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
  1852
      end
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1853
    | Const (@{const_name "op ="}, _) $ t1 =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1854
      SOME (interpret thy model args (eta_expand t 1))
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1855
    | Const (@{const_name "op ="}, _) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1856
      SOME (interpret thy model args (eta_expand t 2))
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1857
    | Const (@{const_name "op &"}, _) $ t1 $ t2 =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1858
      (* 3-valued logic *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1859
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1860
        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
  1861
        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
  1862
        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
  1863
        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
  1864
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1865
        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
  1866
      end
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1867
    | Const (@{const_name "op &"}, _) $ t1 =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1868
      SOME (interpret thy model args (eta_expand t 1))
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1869
    | Const (@{const_name "op &"}, _) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1870
      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
  1871
      (* 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
  1872
      (* "False & undef":                                          *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1873
      (* SOME (Node [Node [TT, FF], Node [FF, FF]], model, args) *)
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1874
    | Const (@{const_name "op |"}, _) $ t1 $ t2 =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1875
      (* 3-valued logic *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1876
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1877
        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
  1878
        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
  1879
        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
  1880
        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
  1881
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1882
        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
  1883
      end
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1884
    | Const (@{const_name "op |"}, _) $ t1 =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1885
      SOME (interpret thy model args (eta_expand t 1))
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1886
    | Const (@{const_name "op |"}, _) =>
22567
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 2))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1888
      (* 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
  1889
      (* "True | undef":                                           *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1890
      (* SOME (Node [Node [TT, TT], Node [TT, FF]], model, args) *)
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1891
    | Const (@{const_name "op -->"}, _) $ t1 $ t2 =>  (* similar to "==>" (Pure) *)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1892
      (* 3-valued logic *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1893
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1894
        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
  1895
        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
  1896
        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
  1897
        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
  1898
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1899
        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
  1900
      end
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1901
    | Const (@{const_name "op -->"}, _) $ t1 =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1902
      SOME (interpret thy model args (eta_expand t 1))
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  1903
    | Const (@{const_name "op -->"}, _) =>
22567
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 2))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1905
      (* 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
  1906
      (* "False --> undef":                                        *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1907
      (* 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
  1908
    | _ => 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
  1909
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1910
  (* 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
  1911
    (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
  1912
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1913
  (* interprets variables and constants whose type is an IDT (this is        *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1914
  (* relatively easy and merely requires us to compute the size of the IDT); *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1915
  (* constructors of IDTs however are properly interpreted by                *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1916
  (* '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
  1917
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1918
  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
  1919
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1920
    val (typs, terms) = model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1921
    (* 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
  1922
    fun interpret_term (Type (s, Ts)) =
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31737
diff changeset
  1923
      (case Datatype.get_info thy s of
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1924
        SOME info =>  (* inductive datatype *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1925
        let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1926
          (* 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
  1927
          val depth = AList.lookup (op =) typs (Type (s, Ts))
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1928
          (* sanity check: depth must be at least 0 *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1929
          val _ = (case depth of SOME n =>
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1930
            if n<0 then
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1931
              raise REFUTE ("IDT_interpreter", "negative depth")
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1932
            else ()
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1933
            | _ => ())
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1934
        in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1935
          (* 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
  1936
          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
  1937
            (* 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
  1938
            SOME (Leaf [], model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1939
          else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1940
            let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1941
              val index               = #index info
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1942
              val descr               = #descr info
29288
253bcf2a5854 avoid polymorphic equality;
wenzelm
parents: 29272
diff changeset
  1943
              val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1944
              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
  1945
              (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1946
              val _ = if Library.exists (fn d =>
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
  1947
                  case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1948
                then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1949
                  raise REFUTE ("IDT_interpreter",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1950
                    "datatype argument (for type "
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26931
diff changeset
  1951
                    ^ Syntax.string_of_typ_global thy (Type (s, Ts))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1952
                    ^ ") is not a variable")
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  1953
                else ()
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1954
              (* 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
  1955
              (* 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
  1956
              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
  1957
                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
  1958
              (* 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
  1959
              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
  1960
              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
  1961
              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
  1962
              (* 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
  1963
              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
  1964
                #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
  1965
              (* prop_formula list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1966
              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
  1967
              (* interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1968
              val intr     = Leaf fms
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1969
              (* 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
  1970
              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
  1971
                | 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
  1972
                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
  1973
              (* prop_formula *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1974
              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
  1975
            in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1976
              (* 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
  1977
              (* condition                                                  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1978
              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
  1979
                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
  1980
                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
  1981
            end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1982
        end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1983
      | 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
  1984
        NONE)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1985
      | 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
  1986
      NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1987
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1988
    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
  1989
      SOME intr =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1990
      (* return an existing interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1991
      SOME (intr, model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1992
    | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1993
      (case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1994
        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
  1995
      | 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
  1996
      | 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
  1997
      | _            => NONE)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  1998
  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
  1999
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2000
  (* 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
  2001
    (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
  2002
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2003
  (* This function imposes an order on the elements of a datatype fragment  *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2004
  (* as follows: C_i x_1 ... x_n < C_j y_1 ... y_m iff i < j or             *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2005
  (* (x_1, ..., x_n) < (y_1, ..., y_m).  With this order, a constructor is  *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2006
  (* a function C_i that maps some argument indices x_1, ..., x_n to the    *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2007
  (* datatype element given by index C_i x_1 ... x_n.  The idea remains the *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2008
  (* same for recursive datatypes, although the computation of indices gets *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2009
  (* a little tricky.                                                       *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2010
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2011
  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
  2012
  let
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2013
    (* returns a list of canonical representations for terms of the type 'T' *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2014
    (* It would be nice if we could just use 'print' for this, but 'print'   *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2015
    (* for IDTs calls 'IDT_constructor_interpreter' again, and this could    *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2016
    (* lead to infinite recursion when we have (mutually) recursive IDTs.    *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2017
    (* (Term.typ * int) list -> Term.typ -> Term.term list *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2018
    fun canonical_terms typs T =
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2019
      (case T of
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2020
        Type ("fun", [T1, T2]) =>
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2021
        (* 'T2' might contain a recursive IDT, so we cannot use 'print' (at *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2022
        (* least not for 'T2'                                               *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2023
        let
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2024
          (* returns a list of lists, each one consisting of n (possibly *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2025
          (* identical) elements from 'xs'                               *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2026
          (* int -> 'a list -> 'a list list *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2027
          fun pick_all 1 xs =
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2028
            map single xs
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2029
          | pick_all n xs =
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2030
            let val rec_pick = pick_all (n-1) xs in
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32950
diff changeset
  2031
              maps (fn x => map (cons x) rec_pick) xs
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2032
            end
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2033
          (* ["x1", ..., "xn"] *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2034
          val terms1 = canonical_terms typs T1
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2035
          (* ["y1", ..., "ym"] *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2036
          val terms2 = canonical_terms typs T2
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2037
          (* [[("x1", "y1"), ..., ("xn", "y1")], ..., *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2038
          (*   [("x1", "ym"), ..., ("xn", "ym")]]     *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2039
          val functions = map (curry (op ~~) terms1)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2040
            (pick_all (length terms1) terms2)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2041
          (* [["(x1, y1)", ..., "(xn, y1)"], ..., *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2042
          (*   ["(x1, ym)", ..., "(xn, ym)"]]     *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2043
          val pairss = map (map HOLogic.mk_prod) functions
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2044
          (* Term.typ *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2045
          val HOLogic_prodT = HOLogic.mk_prodT (T1, T2)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2046
          val HOLogic_setT  = HOLogic.mk_setT HOLogic_prodT
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2047
          (* Term.term *)
36130
9a672f7d488d commented out unsound "lfp"/"gfp" handling + fixed set output syntax;
blanchet
parents: 35746
diff changeset
  2048
          val HOLogic_empty_set = Const (@{const_abbrev Set.empty}, HOLogic_setT)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2049
          val HOLogic_insert    =
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  2050
            Const (@{const_name insert}, HOLogic_prodT --> HOLogic_setT --> HOLogic_setT)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2051
        in
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2052
          (* functions as graphs, i.e. as a (HOL) set of pairs "(x, y)" *)
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
  2053
          map (fn ps => fold_rev (fn pair => fn acc => HOLogic_insert $ pair $ acc) ps
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2054
            HOLogic_empty_set) pairss
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2055
        end
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2056
      | Type (s, Ts) =>
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31737
diff changeset
  2057
        (case Datatype.get_info thy s of
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2058
          SOME info =>
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2059
          (case AList.lookup (op =) typs T of
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2060
            SOME 0 =>
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2061
            (* termination condition to avoid infinite recursion *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2062
            []  (* at depth 0, every IDT is empty *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2063
          | _ =>
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2064
            let
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2065
              val index               = #index info
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2066
              val descr               = #descr info
29288
253bcf2a5854 avoid polymorphic equality;
wenzelm
parents: 29272
diff changeset
  2067
              val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2068
              val typ_assoc           = dtyps ~~ Ts
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2069
              (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2070
              val _ = if Library.exists (fn d =>
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
  2071
                  case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2072
                then
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2073
                  raise REFUTE ("IDT_constructor_interpreter",
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2074
                    "datatype argument (for type "
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26931
diff changeset
  2075
                    ^ Syntax.string_of_typ_global thy T
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2076
                    ^ ") is not a variable")
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2077
                else ()
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2078
              (* decrement depth for the IDT 'T' *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2079
              val typs' = (case AList.lookup (op =) typs T of NONE => typs
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2080
                | SOME n => AList.update (op =) (T, n-1) typs)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2081
              fun constructor_terms terms [] = terms
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2082
                | constructor_terms terms (d::ds) =
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2083
                let
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2084
                  val dT = typ_of_dtyp descr typ_assoc d
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2085
                  val d_terms = canonical_terms typs' dT
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2086
                in
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2087
                  (* C_i x_1 ... x_n < C_i y_1 ... y_n if *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2088
                  (* (x_1, ..., x_n) < (y_1, ..., y_n)    *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2089
                  constructor_terms
25538
58e8ba3b792b map_product and fold_product
haftmann
parents: 25032
diff changeset
  2090
                    (map_product (curry op $) terms d_terms) ds
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2091
                end
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2092
            in
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2093
              (* C_i ... < C_j ... if i < j *)
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32950
diff changeset
  2094
              maps (fn (cname, ctyps) =>
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2095
                let
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2096
                  val cTerm = Const (cname,
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2097
                    map (typ_of_dtyp descr typ_assoc) ctyps ---> T)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2098
                in
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2099
                  constructor_terms [cTerm] ctyps
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32950
diff changeset
  2100
                end) constrs
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2101
            end)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2102
        | NONE =>
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2103
          (* not an inductive datatype; in this case the argument types in *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2104
          (* 'Ts' may not be IDTs either, so 'print' should be safe        *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2105
          map (fn intr => print thy (typs, []) T intr (K false))
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2106
            (make_constants thy (typs, []) T))
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2107
      | _ =>  (* TFree ..., TVar ... *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2108
        map (fn intr => print thy (typs, []) T intr (K false))
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2109
          (make_constants thy (typs, []) T))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2110
    val (typs, terms) = model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2111
  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2112
    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
  2113
      SOME intr =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2114
      (* return an existing interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2115
      SOME (intr, model, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2116
    | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2117
      (case t of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2118
        Const (s, T) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2119
        (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
  2120
          Type (s', Ts') =>
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31737
diff changeset
  2121
          (case Datatype.get_info thy s' of
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2122
            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
  2123
            let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2124
              val index               = #index info
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2125
              val descr               = #descr info
29288
253bcf2a5854 avoid polymorphic equality;
wenzelm
parents: 29272
diff changeset
  2126
              val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2127
              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
  2128
              (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2129
              val _ = if Library.exists (fn d =>
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
  2130
                  case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2131
                then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2132
                  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
  2133
                    "datatype argument (for type "
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26931
diff changeset
  2134
                    ^ Syntax.string_of_typ_global thy (Type (s', Ts'))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2135
                    ^ ") is not a variable")
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2136
                else ()
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2137
              (* 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
  2138
              (* 'Const (s, T)'                                          *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2139
              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
  2140
                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
  2141
                  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
  2142
                    ---> Type (s', Ts')))) constrs
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2143
            in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2144
              case constrs2 of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2145
                [] =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2146
                (* '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
  2147
                NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2148
              | (_, ctypes)::cs =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2149
                let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2150
                  (* 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
  2151
                  (*               depth                                    *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2152
                  val depth = AList.lookup (op =) typs (Type (s', Ts'))
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2153
                  (* this should never happen: at depth 0, this IDT fragment *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2154
                  (* is definitely empty, and in this case we don't need to  *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2155
                  (* interpret its constructors                              *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2156
                  val _ = (case depth of SOME 0 =>
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2157
                      raise REFUTE ("IDT_constructor_interpreter",
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2158
                        "depth is 0")
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2159
                    | _ => ())
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2160
                  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
  2161
                    AList.update (op =) (Type (s', Ts'), n-1) typs)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2162
                  (* elements of the datatype come before elements generated *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2163
                  (* by 'Const (s, T)' iff they are generated by a           *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2164
                  (* constructor in constrs1                                 *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2165
                  val offset = size_of_dtyp thy typs' descr typ_assoc constrs1
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2166
                  (* compute the total (current) size of the datatype *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2167
                  val total = offset +
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2168
                    size_of_dtyp thy typs' descr typ_assoc constrs2
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2169
                  (* sanity check *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2170
                  val _ = if total <> size_of_type thy (typs, [])
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2171
                    (Type (s', Ts')) then
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2172
                      raise REFUTE ("IDT_constructor_interpreter",
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2173
                        "total is not equal to current size")
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2174
                    else ()
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2175
                  (* returns an interpretation where everything is mapped to *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2176
                  (* an "undefined" element of the datatype                  *)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2177
                  fun make_undef [] =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2178
                    Leaf (replicate total False)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2179
                    | make_undef (d::ds) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2180
                    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2181
                      (* compute the current size of the type 'd' *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2182
                      val dT   = typ_of_dtyp descr typ_assoc d
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2183
                      val size = size_of_type thy (typs, []) dT
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2184
                    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2185
                      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
  2186
                    end
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2187
                  (* returns the interpretation for a constructor *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2188
                  fun make_constr [] offset =
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2189
                    if offset < total then
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2190
                      (Leaf (replicate offset False @ True ::
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2191
                        (replicate (total - offset - 1) False)), offset + 1)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2192
                    else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2193
                      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
  2194
                        "offset >= total")
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2195
                    | make_constr (d::ds) offset =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2196
                    let
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2197
                      (* Term.typ *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2198
                      val dT = typ_of_dtyp descr typ_assoc d
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2199
                      (* compute canonical term representations for all   *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2200
                      (* elements of the type 'd' (with the reduced depth *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2201
                      (* for the IDT)                                     *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2202
                      val terms' = canonical_terms typs' dT
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2203
                      (* sanity check *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2204
                      val _ =
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2205
                        if length terms' <> size_of_type thy (typs', []) dT
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2206
                        then
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2207
                          raise REFUTE ("IDT_constructor_interpreter",
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2208
                            "length of terms' is not equal to old size")
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2209
                        else ()
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2210
                      (* compute canonical term representations for all   *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2211
                      (* elements of the type 'd' (with the current depth *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2212
                      (* for the IDT)                                     *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2213
                      val terms = canonical_terms typs dT
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2214
                      (* sanity check *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2215
                      val _ =
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2216
                        if length terms <> size_of_type thy (typs, []) dT
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2217
                        then
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2218
                          raise REFUTE ("IDT_constructor_interpreter",
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2219
                            "length of terms is not equal to current size")
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2220
                        else ()
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2221
                      (* sanity check *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2222
                      val _ =
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2223
                        if length terms < length terms' then
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2224
                          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
  2225
                            "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
  2226
                        else ()
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2227
                      (* sanity check: every element of terms' must also be *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2228
                      (*               present in terms                     *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2229
                      val _ =
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2230
                        if List.all (member (op =) terms) terms' then ()
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2231
                        else
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2232
                          raise REFUTE ("IDT_constructor_interpreter",
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2233
                            "element has disappeared")
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2234
                      (* sanity check: the order on elements of terms' is    *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2235
                      (*               the same in terms, for those elements *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2236
                      val _ =
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2237
                        let
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2238
                          fun search (x::xs) (y::ys) =
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2239
                                if x = y then search xs ys else search (x::xs) ys
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2240
                            | search (x::xs) [] =
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2241
                                raise REFUTE ("IDT_constructor_interpreter",
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2242
                                  "element order not preserved")
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2243
                            | search [] _ = ()
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2244
                        in  search terms' terms  end
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2245
                      (* int * interpretation list *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2246
                      val (intrs, new_offset) =
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2247
                        fold_map (fn t_elem => fn off =>
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2248
                          (* if 't_elem' existed at the previous depth,    *)
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2249
                          (* proceed recursively, otherwise map the entire *)
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2250
                          (* subtree to "undefined"                        *)
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2251
                          if t_elem mem terms' then
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2252
                            make_constr ds off
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2253
                          else
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2254
                            (make_undef ds, off))
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2255
                        terms offset
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2256
                    in
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2257
                      (Node intrs, new_offset)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2258
                    end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2259
                in
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2260
                  SOME (fst (make_constr ctypes offset), model, args)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2261
                end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2262
            end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2263
          | 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
  2264
            NONE)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2265
        | _ =>  (* 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
  2266
          NONE)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2267
      | _ =>  (* 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
  2268
        NONE)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2269
  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
  2270
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2271
  (* 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
  2272
    (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
  2273
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2274
  (* 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
  2275
  (* '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
  2276
  (* 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
  2277
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2278
  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
  2279
    (* 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
  2280
    (* 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
  2281
    (* 't'                                                                 *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2282
    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
  2283
      (Const (s, T), params) =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2284
      (* 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
  2285
      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
  2286
        case result of
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2287
          SOME _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2288
          result  (* just keep 'result' *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2289
        | NONE =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2290
          if member (op =) (#rec_names info) s then
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2291
            (* we do have a recursion operator of one of the (mutually *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2292
            (* recursive) datatypes given by 'info'                    *)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2293
            let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2294
              (* 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
  2295
              (* (mutually recursive) datatypes within the same descriptor *)
33002
f3f02f36a3e2 uniform use of Integer.add/mult/sum/prod;
wenzelm
parents: 32952
diff changeset
  2296
              val mconstrs_count =
f3f02f36a3e2 uniform use of Integer.add/mult/sum/prod;
wenzelm
parents: 32952
diff changeset
  2297
                Integer.sum (map (fn (_, (_, _, cs)) => length cs) (#descr info))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2298
            in
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2299
              if mconstrs_count < length params then
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2300
                (* 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
  2301
                (* '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
  2302
                NONE
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2303
              else if mconstrs_count > length params then
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2304
                (* 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
  2305
                (* 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
  2306
                (* 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
  2307
                (* 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
  2308
                (* course).                                                 *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2309
                SOME (interpret thy model args (eta_expand t
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2310
                  (mconstrs_count - length params)))
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2311
              else  (* mconstrs_count = length params *)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2312
                let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2313
                  (* interpret each parameter separately *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2314
                  val (p_intrs, (model', args')) = fold_map (fn p => fn (m, a) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2315
                    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2316
                      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
  2317
                    in
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2318
                      (i, (m', a'))
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2319
                    end) params (model, args)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2320
                  val (typs, _) = model'
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2321
                  (* 'index' is /not/ necessarily the index of the IDT that *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2322
                  (* the recursion operator is associated with, but merely  *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2323
                  (* the index of some mutually recursive IDT               *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2324
                  val index         = #index info
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2325
                  val descr         = #descr info
29288
253bcf2a5854 avoid polymorphic equality;
wenzelm
parents: 29272
diff changeset
  2326
                  val (_, dtyps, _) = the (AList.lookup (op =) descr index)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2327
                  (* sanity check: we assume that the order of constructors *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2328
                  (*               in 'descr' is the same as the order of   *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2329
                  (*               corresponding parameters, otherwise the  *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2330
                  (*               association code below won't match the   *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2331
                  (*               right constructors/parameters; we also   *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2332
                  (*               assume that the order of recursion       *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2333
                  (*               operators in '#rec_names info' is the    *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2334
                  (*               same as the order of corresponding       *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2335
                  (*               datatypes in 'descr'                     *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2336
                  val _ = if map fst descr <> (0 upto (length descr - 1)) then
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2337
                      raise REFUTE ("IDT_recursion_interpreter",
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2338
                        "order of constructors and corresponding parameters/" ^
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2339
                          "recursion operators and corresponding datatypes " ^
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2340
                          "different?")
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2341
                    else ()
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2342
                  (* sanity check: every element in 'dtyps' must be a *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2343
                  (*               'DtTFree'                          *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2344
                  val _ = if Library.exists (fn d =>
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
  2345
                    case d of Datatype_Aux.DtTFree _ => false
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2346
                            | _ => true) dtyps
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2347
                    then
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2348
                      raise REFUTE ("IDT_recursion_interpreter",
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2349
                        "datatype argument is not a variable")
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2350
                    else ()
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2351
                  (* the type of a recursion operator is *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2352
                  (* [T1, ..., Tn, IDT] ---> Tresult     *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2353
                  val IDT = List.nth (binder_types T, mconstrs_count)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2354
                  (* by our assumption on the order of recursion operators *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2355
                  (* and datatypes, this is the index of the datatype      *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2356
                  (* corresponding to the given recursion operator         *)
31986
a68f88d264f7 dropped find_index_eq
haftmann
parents: 31943
diff changeset
  2357
                  val idt_index = find_index (fn s' => s' = s) (#rec_names info)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2358
                  (* mutually recursive types must have the same type   *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2359
                  (* parameters, unless the mutual recursion comes from *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2360
                  (* indirect recursion                                 *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2361
                  fun rec_typ_assoc acc [] =
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2362
                    acc
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2363
                    | rec_typ_assoc acc ((d, T)::xs) =
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2364
                    (case AList.lookup op= acc d of
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2365
                      NONE =>
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2366
                      (case d of
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
  2367
                        Datatype_Aux.DtTFree _ =>
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2368
                        (* add the association, proceed *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2369
                        rec_typ_assoc ((d, T)::acc) xs
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
  2370
                      | Datatype_Aux.DtType (s, ds) =>
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2371
                        let
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2372
                          val (s', Ts) = dest_Type T
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2373
                        in
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2374
                          if s=s' then
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2375
                            rec_typ_assoc ((d, T)::acc) ((ds ~~ Ts) @ xs)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2376
                          else
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2377
                            raise REFUTE ("IDT_recursion_interpreter",
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2378
                              "DtType/Type mismatch")
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2379
                        end
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
  2380
                      | Datatype_Aux.DtRec i =>
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2381
                        let
29288
253bcf2a5854 avoid polymorphic equality;
wenzelm
parents: 29272
diff changeset
  2382
                          val (_, ds, _) = the (AList.lookup (op =) descr i)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2383
                          val (_, Ts)    = dest_Type T
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2384
                        in
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2385
                          rec_typ_assoc ((d, T)::acc) ((ds ~~ Ts) @ xs)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2386
                        end)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2387
                    | SOME T' =>
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2388
                      if T=T' then
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2389
                        (* ignore the association since it's already *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2390
                        (* present, proceed                          *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2391
                        rec_typ_assoc acc xs
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2392
                      else
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2393
                        raise REFUTE ("IDT_recursion_interpreter",
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2394
                          "different type associations for the same dtyp"))
29288
253bcf2a5854 avoid polymorphic equality;
wenzelm
parents: 29272
diff changeset
  2395
                  val typ_assoc = filter
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
  2396
                    (fn (Datatype_Aux.DtTFree _, _) => true | (_, _) => false)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2397
                    (rec_typ_assoc []
29288
253bcf2a5854 avoid polymorphic equality;
wenzelm
parents: 29272
diff changeset
  2398
                      (#2 (the (AList.lookup (op =) descr idt_index)) ~~ (snd o dest_Type) IDT))
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2399
                  (* sanity check: typ_assoc must associate types to the   *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2400
                  (*               elements of 'dtyps' (and only to those) *)
33038
8f9594c31de4 dropped redundant gen_ prefix
haftmann
parents: 33037
diff changeset
  2401
                  val _ = if not (eq_set (op =) (dtyps, map fst typ_assoc))
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2402
                    then
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2403
                      raise REFUTE ("IDT_recursion_interpreter",
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2404
                        "type association has extra/missing elements")
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2405
                    else ()
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2406
                  (* 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
  2407
                  (* 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
  2408
                  (* (int * interpretation list) list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2409
                  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
  2410
                    let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2411
                      val c_return_typ = typ_of_dtyp descr typ_assoc
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
  2412
                        (Datatype_Aux.DtRec idx)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2413
                    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2414
                      (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
  2415
                        (#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
  2416
                          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
  2417
                          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
  2418
                          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
  2419
                    end) descr
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2420
                  (* associate constructors with corresponding parameters *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2421
                  (* (int * (interpretation * interpretation) list) list *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2422
                  val (mc_p_intrs, p_intrs') = fold_map
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2423
                    (fn (idx, c_intrs) => fn p_intrs' =>
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2424
                      let
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2425
                        val len = length c_intrs
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2426
                      in
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2427
                        ((idx, c_intrs ~~ List.take (p_intrs', len)),
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2428
                          List.drop (p_intrs', len))
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2429
                      end) mc_intrs p_intrs
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2430
                  (* sanity check: no 'p_intr' may be left afterwards *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2431
                  val _ = if p_intrs' <> [] then
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2432
                      raise REFUTE ("IDT_recursion_interpreter",
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2433
                        "more parameter than constructor interpretations")
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2434
                    else ()
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2435
                  (* The recursion operator, applied to 'mconstrs_count'     *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2436
                  (* arguments, is a function that maps every element of the *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2437
                  (* inductive datatype to an element of some result type.   *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2438
                  (* Recursion operators for mutually recursive IDTs are     *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2439
                  (* translated simultaneously.                              *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2440
                  (* Since the order on datatype elements is given by an     *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2441
                  (* order on constructors (and then by the order on         *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2442
                  (* argument tuples), we can simply copy corresponding      *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2443
                  (* subtrees from 'p_intrs', in the order in which they are *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2444
                  (* given.                                                  *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2445
                  (* interpretation * interpretation -> interpretation list *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2446
                  fun ci_pi (Leaf xs, pi) =
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2447
                    (* if the constructor does not match the arguments to a *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2448
                    (* defined element of the IDT, the corresponding value  *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2449
                    (* of the parameter must be ignored                     *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2450
                    if List.exists (equal True) xs then [pi] else []
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2451
                    | ci_pi (Node xs, Node ys) =
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32950
diff changeset
  2452
                    maps ci_pi (xs ~~ ys)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2453
                    | ci_pi (Node _, Leaf _) =
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2454
                    raise REFUTE ("IDT_recursion_interpreter",
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2455
                      "constructor takes more arguments than the " ^
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2456
                        "associated parameter")
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2457
                  (* (int * interpretation list) list *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2458
                  val rec_operators = map (fn (idx, c_p_intrs) =>
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32950
diff changeset
  2459
                    (idx, maps ci_pi c_p_intrs)) mc_p_intrs
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2460
                  (* sanity check: every recursion operator must provide as  *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2461
                  (*               many values as the corresponding datatype *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2462
                  (*               has elements                              *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2463
                  val _ = map (fn (idx, intrs) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2464
                    let
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2465
                      val T = typ_of_dtyp descr typ_assoc
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
  2466
                        (Datatype_Aux.DtRec idx)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2467
                    in
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2468
                      if length intrs <> size_of_type thy (typs, []) T then
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2469
                        raise REFUTE ("IDT_recursion_interpreter",
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2470
                          "wrong number of interpretations for rec. operator")
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2471
                      else ()
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2472
                    end) rec_operators
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2473
                  (* For non-recursive datatypes, we are pretty much done at *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2474
                  (* this point.  For recursive datatypes however, we still  *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2475
                  (* need to apply the interpretations in 'rec_operators' to *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2476
                  (* (recursively obtained) interpretations for recursive    *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2477
                  (* constructor arguments.  To do so more efficiently, we   *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2478
                  (* copy 'rec_operators' into arrays first.  Each Boolean   *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2479
                  (* indicates whether the recursive arguments have been     *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2480
                  (* considered already.                                     *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2481
                  (* (int * (bool * interpretation) Array.array) list *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2482
                  val REC_OPERATORS = map (fn (idx, intrs) =>
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2483
                    (idx, Array.fromList (map (pair false) intrs)))
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2484
                    rec_operators
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2485
                  (* 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
  2486
                  (* 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
  2487
                  (* 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
  2488
                  (* returned                                              *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2489
                  (* 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
  2490
                  fun get_args (Leaf xs) elem =
31986
a68f88d264f7 dropped find_index_eq
haftmann
parents: 31943
diff changeset
  2491
                    if find_index (fn x => x = True) xs = elem then
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2492
                      SOME []
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2493
                    else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2494
                      NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2495
                    | 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
  2496
                    let
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2497
                      (* interpretation list * int -> int list option *)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2498
                      fun search ([], _) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2499
                        NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2500
                        | search (x::xs, n) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2501
                        (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
  2502
                          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
  2503
                        | 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
  2504
                    in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2505
                      search (xs, 0)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2506
                    end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2507
                  (* 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
  2508
                  (* 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
  2509
                  (* 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
  2510
                  (* 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
  2511
                  fun get_cargs idx elem =
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2512
                  let
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2513
                    (* int * interpretation list -> int * int list *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2514
                    fun get_cargs_rec (_, []) =
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2515
                      raise REFUTE ("IDT_recursion_interpreter",
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2516
                        "no matching constructor found for datatype element")
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2517
                      | get_cargs_rec (n, x::xs) =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2518
                        (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
  2519
                          SOME args => (n, args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2520
                        | 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
  2521
                    in
29288
253bcf2a5854 avoid polymorphic equality;
wenzelm
parents: 29272
diff changeset
  2522
                      get_cargs_rec (0, the (AList.lookup (op =) mc_intrs idx))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2523
                    end
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2524
                  (* computes one entry in 'REC_OPERATORS', and recursively *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2525
                  (* all entries needed for it, where 'idx' gives the       *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2526
                  (* datatype and 'elem' the element of it                  *)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2527
                  (* int -> int -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2528
                  fun compute_array_entry idx elem =
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2529
                  let
29288
253bcf2a5854 avoid polymorphic equality;
wenzelm
parents: 29272
diff changeset
  2530
                    val arr          = the (AList.lookup (op =) REC_OPERATORS idx)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2531
                    val (flag, intr) = Array.sub (arr, elem)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2532
                  in
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2533
                    if flag then
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2534
                      (* simply return the previously computed result *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2535
                      intr
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2536
                    else
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2537
                      (* we have to apply 'intr' to interpretations for all *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2538
                      (* recursive arguments                                *)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2539
                      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2540
                        (* int * int list *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2541
                        val (c, args) = get_cargs idx elem
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2542
                        (* find the indices of the constructor's /recursive/ *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2543
                        (* arguments                                         *)
29288
253bcf2a5854 avoid polymorphic equality;
wenzelm
parents: 29272
diff changeset
  2544
                        val (_, _, constrs) = the (AList.lookup (op =) descr idx)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2545
                        val (_, dtyps)      = List.nth (constrs, c)
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 33246
diff changeset
  2546
                        val rec_dtyps_args  = filter
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
  2547
                          (Datatype_Aux.is_rec_type o fst) (dtyps ~~ args)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2548
                        (* map those indices to interpretations *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2549
                        val rec_dtyps_intrs = map (fn (dtyp, arg) =>
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2550
                          let
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2551
                            val dT     = typ_of_dtyp descr typ_assoc dtyp
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2552
                            val consts = make_constants thy (typs, []) dT
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2553
                            val arg_i  = List.nth (consts, arg)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2554
                          in
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2555
                            (dtyp, arg_i)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2556
                          end) rec_dtyps_args
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2557
                        (* takes the dtyp and interpretation of an element, *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2558
                        (* and computes the interpretation for the          *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2559
                        (* corresponding recursive argument                 *)
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
  2560
                        fun rec_intr (Datatype_Aux.DtRec i) (Leaf xs) =
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2561
                          (* recursive argument is "rec_i params elem" *)
31986
a68f88d264f7 dropped find_index_eq
haftmann
parents: 31943
diff changeset
  2562
                          compute_array_entry i (find_index (fn x => x = True) xs)
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
  2563
                          | rec_intr (Datatype_Aux.DtRec _) (Node _) =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2564
                          raise REFUTE ("IDT_recursion_interpreter",
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2565
                            "interpretation for IDT is a node")
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
  2566
                          | rec_intr (Datatype_Aux.DtType ("fun", [dt1, dt2]))
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2567
                            (Node xs) =
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2568
                          (* recursive argument is something like     *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2569
                          (* "\<lambda>x::dt1. rec_? params (elem x)" *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2570
                          Node (map (rec_intr dt2) xs)
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
  2571
                          | rec_intr (Datatype_Aux.DtType ("fun", [_, _]))
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2572
                            (Leaf _) =
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2573
                          raise REFUTE ("IDT_recursion_interpreter",
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2574
                            "interpretation for function dtyp is a leaf")
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2575
                          | rec_intr _ _ =
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2576
                          (* admissibility ensures that every recursive type *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2577
                          (* is of the form 'Dt_1 -> ... -> Dt_k ->          *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2578
                          (* (DtRec i)'                                      *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2579
                          raise REFUTE ("IDT_recursion_interpreter",
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2580
                            "non-recursive codomain in recursive dtyp")
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2581
                        (* obtain interpretations for recursive arguments *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2582
                        (* interpretation list *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2583
                        val arg_intrs = map (uncurry rec_intr) rec_dtyps_intrs
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2584
                        (* apply 'intr' to all recursive arguments *)
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
  2585
                        val result = fold (fn arg_i => fn i =>
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
  2586
                          interpretation_apply (i, arg_i)) arg_intrs intr
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2587
                        (* update 'REC_OPERATORS' *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2588
                        val _ = Array.update (arr, elem, (true, result))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2589
                      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2590
                        result
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2591
                      end
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2592
                  end
29288
253bcf2a5854 avoid polymorphic equality;
wenzelm
parents: 29272
diff changeset
  2593
                  val idt_size = Array.length (the (AList.lookup (op =) REC_OPERATORS idt_index))
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2594
                  (* sanity check: the size of 'IDT' should be 'idt_size' *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2595
                  val _ = if idt_size <> size_of_type thy (typs, []) IDT then
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2596
                        raise REFUTE ("IDT_recursion_interpreter",
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2597
                          "unexpected size of IDT (wrong type associated?)")
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2598
                      else ()
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2599
                  (* interpretation *)
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 33055
diff changeset
  2600
                  val rec_op = Node (map_range (compute_array_entry idt_index) idt_size)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2601
                in
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2602
                  SOME (rec_op, model', args')
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2603
                end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2604
            end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2605
          else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2606
            NONE  (* not a recursion operator of this datatype *)
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31737
diff changeset
  2607
        ) (Datatype.get_all thy) NONE
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2608
    | _ =>  (* 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
  2609
      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
  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 *)
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
  2613
29956
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2614
  fun set_interpreter thy model args t =
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2615
  let
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2616
    val (typs, terms) = model
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2617
  in
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2618
    case AList.lookup (op =) terms t of
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2619
      SOME intr =>
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2620
      (* return an existing interpretation *)
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2621
      SOME (intr, model, args)
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2622
    | NONE =>
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2623
      (case t of
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2624
      (* 'Collect' == identity *)
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2625
        Const (@{const_name Collect}, _) $ t1 =>
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2626
        SOME (interpret thy model args t1)
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2627
      | Const (@{const_name Collect}, _) =>
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2628
        SOME (interpret thy model args (eta_expand t 1))
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2629
      (* 'op :' == application *)
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2630
      | Const (@{const_name "op :"}, _) $ t1 $ t2 =>
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2631
        SOME (interpret thy model args (t2 $ t1))
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2632
      | Const (@{const_name "op :"}, _) $ t1 =>
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2633
        SOME (interpret thy model args (eta_expand t 1))
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2634
      | Const (@{const_name "op :"}, _) =>
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2635
        SOME (interpret thy model args (eta_expand t 2))
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2636
      | _ => NONE)
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2637
  end;
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2638
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2639
  (* theory -> model -> arguments -> Term.term ->
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2640
    (interpretation * model * arguments) option *)
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  2641
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2642
  (* 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
  2643
  (* 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
  2644
  (* 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
  2645
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2646
  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
  2647
    case t of
29820
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
  2648
      Const (@{const_name Finite_Set.card},
30312
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  2649
        Type ("fun", [Type ("fun", [T, Type ("bool", [])]),
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  2650
                      Type ("nat", [])])) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2651
      let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2652
        (* interpretation -> int *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2653
        fun number_of_elements (Node xs) =
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2654
            fold (fn x => fn n =>
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2655
              if x = TT then
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2656
                n + 1
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2657
              else if x = FF then
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2658
                n
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2659
              else
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2660
                raise REFUTE ("Finite_Set_card_interpreter",
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2661
                  "interpretation for set type does not yield a Boolean"))
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2662
              xs 0
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2663
          | number_of_elements (Leaf _) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2664
          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
  2665
            "interpretation for set type is a leaf")
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2666
        val size_of_nat = size_of_type thy model (Type ("nat", []))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2667
        (* takes an interpretation for a set and returns an interpretation *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2668
        (* for a 'nat' denoting the set's cardinality                      *)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2669
        (* interpretation -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2670
        fun card i =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2671
          let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2672
            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
  2673
          in
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2674
            if n<size_of_nat then
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2675
              Leaf ((replicate n False) @ True ::
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2676
                (replicate (size_of_nat-n-1) False))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2677
            else
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2678
              Leaf (replicate size_of_nat False)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2679
          end
30312
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  2680
        val set_constants =
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  2681
          make_constants thy model (Type ("fun", [T, Type ("bool", [])]))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2682
      in
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2683
        SOME (Node (map card set_constants), model, args)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2684
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2685
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2686
      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
  2687
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2688
  (* 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
  2689
    (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
  2690
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2691
  (* 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
  2692
  (* 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
  2693
  (* below is more efficient                                                *)
22093
98e3e9f00192 interpreter for Finite_Set.finite added
webertj
parents: 22092
diff changeset
  2694
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2695
  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
  2696
    case t of
29820
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
  2697
      Const (@{const_name Finite_Set.finite},
30312
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  2698
        Type ("fun", [Type ("fun", [T, Type ("bool", [])]),
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  2699
                      Type ("bool", [])])) $ _ =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2700
        (* 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
  2701
        (* "finite"                                                  *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2702
        SOME (TT, model, args)
29820
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
  2703
    | Const (@{const_name Finite_Set.finite},
30312
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  2704
        Type ("fun", [Type ("fun", [T, Type ("bool", [])]),
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  2705
                      Type ("bool", [])])) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2706
      let
30312
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  2707
        val size_of_set =
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  2708
          size_of_type thy model (Type ("fun", [T, Type ("bool", [])]))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2709
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2710
        (* 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
  2711
        (* "finite"                                                  *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2712
        SOME (Node (replicate size_of_set TT), model, args)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2713
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2714
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2715
      NONE;
22093
98e3e9f00192 interpreter for Finite_Set.finite added
webertj
parents: 22092
diff changeset
  2716
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2717
  (* 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
  2718
    (interpretation * model * arguments) option *)
22093
98e3e9f00192 interpreter for Finite_Set.finite added
webertj
parents: 22092
diff changeset
  2719
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 34120
diff changeset
  2720
  (* only an optimization: 'less' could in principle be interpreted with *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2721
  (* interpreters available already (using its definition), but the code     *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2722
  (* 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
  2723
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2724
  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
  2725
    case t of
35092
cfe605c54e50 moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents: 34974
diff changeset
  2726
      Const (@{const_name Orderings.less}, Type ("fun", [Type ("nat", []),
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2727
        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
  2728
      let
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2729
        val size_of_nat = size_of_type thy model (Type ("nat", []))
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2730
        (* the 'n'-th nat is not less than the first 'n' nats, while it *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2731
        (* is less than the remaining 'size_of_nat - n' nats            *)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2732
        (* int -> interpretation *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2733
        fun less n = Node ((replicate n FF) @ (replicate (size_of_nat - n) TT))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2734
      in
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2735
        SOME (Node (map less (1 upto size_of_nat)), model, args)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2736
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2737
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2738
      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
  2739
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2740
  (* 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
  2741
    (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
  2742
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 34120
diff changeset
  2743
  (* only an optimization: 'plus' could in principle be interpreted with *)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2744
  (* 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
  2745
  (* 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
  2746
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2747
  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
  2748
    case t of
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35092
diff changeset
  2749
      Const (@{const_name Groups.plus}, Type ("fun", [Type ("nat", []),
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2750
        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
  2751
      let
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2752
        val size_of_nat = size_of_type thy model (Type ("nat", []))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2753
        (* int -> int -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2754
        fun plus m n =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2755
          let
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2756
            val element = m + n
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2757
          in
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2758
            if element > size_of_nat - 1 then
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2759
              Leaf (replicate size_of_nat False)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2760
            else
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2761
              Leaf ((replicate element False) @ True ::
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2762
                (replicate (size_of_nat - element - 1) False))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2763
          end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2764
      in
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 33055
diff changeset
  2765
        SOME (Node (map_range (fn m => Node (map_range (plus m) size_of_nat)) size_of_nat),
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 33055
diff changeset
  2766
          model, args)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2767
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2768
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2769
      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
  2770
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2771
  (* 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
  2772
    (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
  2773
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 34120
diff changeset
  2774
  (* only an optimization: 'minus' could in principle be interpreted *)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2775
  (* 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
  2776
  (* 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
  2777
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2778
  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
  2779
    case t of
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35092
diff changeset
  2780
      Const (@{const_name Groups.minus}, Type ("fun", [Type ("nat", []),
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2781
        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
  2782
      let
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2783
        val size_of_nat = size_of_type thy model (Type ("nat", []))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2784
        (* int -> int -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2785
        fun minus m n =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2786
          let
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2787
            val element = Int.max (m-n, 0)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2788
          in
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2789
            Leaf ((replicate element False) @ True ::
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2790
              (replicate (size_of_nat - element - 1) False))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2791
          end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2792
      in
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 33055
diff changeset
  2793
        SOME (Node (map_range (fn m => Node (map_range (minus m) size_of_nat)) size_of_nat),
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 33055
diff changeset
  2794
          model, args)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2795
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2796
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2797
      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
  2798
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2799
  (* 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
  2800
    (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
  2801
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 34120
diff changeset
  2802
  (* only an optimization: 'times' could in principle be interpreted *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2803
  (* with interpreters available already (using its definition), but the *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2804
  (* 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
  2805
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2806
  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
  2807
    case t of
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35092
diff changeset
  2808
      Const (@{const_name Groups.times}, Type ("fun", [Type ("nat", []),
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2809
        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
  2810
      let
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2811
        val size_of_nat = size_of_type thy model (Type ("nat", []))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2812
        (* nat -> nat -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2813
        fun mult m n =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2814
          let
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2815
            val element = m * n
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2816
          in
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2817
            if element > size_of_nat - 1 then
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2818
              Leaf (replicate size_of_nat False)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2819
            else
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2820
              Leaf ((replicate element False) @ True ::
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2821
                (replicate (size_of_nat - element - 1) False))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2822
          end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2823
      in
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 33055
diff changeset
  2824
        SOME (Node (map_range (fn m => Node (map_range (mult m) size_of_nat)) size_of_nat),
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 33055
diff changeset
  2825
          model, args)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2826
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2827
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2828
      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
  2829
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2830
  (* 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
  2831
    (interpretation * model * arguments) option *)
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15611
diff changeset
  2832
23029
79ee75dc1e59 constant op @ now named append
haftmann
parents: 22997
diff changeset
  2833
  (* only an optimization: 'append' could in principle be interpreted with *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2834
  (* interpreters available already (using its definition), but the code   *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2835
  (* below is more efficient                                               *)
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15611
diff changeset
  2836
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2837
  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
  2838
    case t of
29820
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
  2839
      Const (@{const_name List.append}, Type ("fun", [Type ("List.list", [T]), Type ("fun",
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2840
        [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
  2841
      let
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2842
        val size_elem   = size_of_type thy model T
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2843
        val size_list   = size_of_type thy model (Type ("List.list", [T]))
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2844
        (* maximal length of lists; 0 if we only consider the empty list *)
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2845
        val list_length = let
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2846
            (* int -> int -> int -> int *)
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2847
            fun list_length_acc len lists total =
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2848
              if lists = total then
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2849
                len
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2850
              else if lists < total then
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2851
                list_length_acc (len+1) (lists*size_elem) (total-lists)
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2852
              else
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2853
                raise REFUTE ("List_append_interpreter",
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2854
                  "size_list not equal to 1 + size_elem + ... + " ^
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2855
                    "size_elem^len, for some len")
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2856
          in
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2857
            list_length_acc 0 1 size_list
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2858
          end
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2859
        val elements = 0 upto (size_list-1)
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2860
        (* FIXME: there should be a nice formula, which computes the same as *)
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2861
        (*        the following, but without all this intermediate tree      *)
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2862
        (*        length/offset stuff                                        *)
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2863
        (* associate each list with its length and offset in a complete tree *)
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2864
        (* of width 'size_elem' and depth 'length_list' (with 'size_list'    *)
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2865
        (* nodes total)                                                      *)
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2866
        (* (int * (int * int)) list *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2867
        val (lenoff_lists, _) = fold_map (fn elem => fn (offsets, off) =>
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2868
          (* corresponds to a pre-order traversal of the tree *)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2869
          let
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2870
            val len = length offsets
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2871
            (* associate the given element with len/off *)
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2872
            val assoc = (elem, (len, off))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2873
          in
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2874
            if len < list_length then
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2875
              (* go to first child node *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2876
              (assoc, (off :: offsets, off * size_elem))
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2877
            else if off mod size_elem < size_elem - 1 then
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2878
              (* go to next sibling node *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2879
              (assoc, (offsets, off + 1))
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2880
            else
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2881
              (* go back up the stack until we find a level where we can go *)
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2882
              (* to the next sibling node                                   *)
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2883
              let
33955
fff6f11b1f09 curried take/drop
haftmann
parents: 33522
diff changeset
  2884
                val offsets' = dropwhile
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2885
                  (fn off' => off' mod size_elem = size_elem - 1) offsets
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2886
              in
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2887
                case offsets' of
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2888
                  [] =>
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2889
                  (* we're at the last node in the tree; the next value *)
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2890
                  (* won't be used anyway                               *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2891
                  (assoc, ([], 0))
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2892
                | off'::offs' =>
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2893
                  (* go to next sibling node *)
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2894
                  (assoc, (offs', off' + 1))
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2895
              end
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  2896
          end) elements ([], 0)
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2897
        (* we also need the reverse association (from length/offset to *)
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2898
        (* index)                                                      *)
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2899
        val lenoff'_lists = map Library.swap lenoff_lists
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2900
        (* returns the interpretation for "(list no. m) @ (list no. n)" *)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2901
        (* nat -> nat -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2902
        fun append m n =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2903
          let
29288
253bcf2a5854 avoid polymorphic equality;
wenzelm
parents: 29272
diff changeset
  2904
            val (len_m, off_m) = the (AList.lookup (op =) lenoff_lists m)
253bcf2a5854 avoid polymorphic equality;
wenzelm
parents: 29272
diff changeset
  2905
            val (len_n, off_n) = the (AList.lookup (op =) lenoff_lists n)
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2906
            val len_elem = len_m + len_n
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2907
            val off_elem = off_m * power (size_elem, len_n) + off_n
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2908
          in
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2909
            case AList.lookup op= lenoff'_lists (len_elem, off_elem)  of
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2910
              NONE =>
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2911
              (* undefined *)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2912
              Leaf (replicate size_list False)
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2913
            | SOME element =>
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2914
              Leaf ((replicate element False) @ True ::
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2915
                (replicate (size_list - element - 1) False))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2916
          end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2917
      in
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2918
        SOME (Node (map (fn m => Node (map (append m) elements)) elements),
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  2919
          model, args)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2920
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2921
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2922
      NONE;
15767
8ed9fcc004fe support for recursion over mutually recursive IDTs
webertj
parents: 15611
diff changeset
  2923
36130
9a672f7d488d commented out unsound "lfp"/"gfp" handling + fixed set output syntax;
blanchet
parents: 35746
diff changeset
  2924
(* UNSOUND
9a672f7d488d commented out unsound "lfp"/"gfp" handling + fixed set output syntax;
blanchet
parents: 35746
diff changeset
  2925
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2926
  (* 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
  2927
    (interpretation * model * arguments) option *)
16050
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15794
diff changeset
  2928
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2929
  (* 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
  2930
  (* 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
  2931
  (* below is more efficient                                             *)
16050
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15794
diff changeset
  2932
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  2933
  fun lfp_interpreter thy model args t =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2934
    case t of
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  2935
      Const (@{const_name lfp}, Type ("fun", [Type ("fun",
30312
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  2936
        [Type ("fun", [T, Type ("bool", [])]),
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  2937
         Type ("fun", [_, Type ("bool", [])])]),
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  2938
         Type ("fun", [_, Type ("bool", [])])])) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2939
      let
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2940
        val size_elem = size_of_type thy model T
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2941
        (* the universe (i.e. the set that contains every element) *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2942
        val i_univ = Node (replicate size_elem TT)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2943
        (* all sets with elements from type 'T' *)
30312
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  2944
        val i_sets =
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  2945
          make_constants thy model (Type ("fun", [T, Type ("bool", [])]))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2946
        (* all functions that map sets to sets *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2947
        val i_funs = make_constants thy model (Type ("fun",
30312
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  2948
          [Type ("fun", [T, Type ("bool", [])]),
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  2949
           Type ("fun", [T, Type ("bool", [])])]))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2950
        (* "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
  2951
        (* interpretation * interpretation -> bool *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2952
        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
  2953
          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
  2954
            (subs ~~ sups)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2955
          | is_subset (_, _) =
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  2956
          raise REFUTE ("lfp_interpreter",
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2957
            "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
  2958
        (* interpretation * interpretation -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2959
        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
  2960
          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
  2961
            (xs ~~ ys))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2962
          | intersection (_, _) =
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  2963
          raise REFUTE ("lfp_interpreter",
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2964
            "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
  2965
        (* interpretation -> interpretaion *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2966
        fun lfp (Node resultsets) =
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
  2967
          fold (fn (set, resultset) => fn acc =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2968
            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
  2969
              intersection (acc, set)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2970
            else
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
  2971
              acc) (i_sets ~~ resultsets) i_univ
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2972
          | lfp _ =
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  2973
            raise REFUTE ("lfp_interpreter",
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2974
              "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
  2975
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2976
        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
  2977
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2978
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2979
      NONE;
16050
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15794
diff changeset
  2980
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2981
  (* 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
  2982
    (interpretation * model * arguments) option *)
16050
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15794
diff changeset
  2983
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2984
  (* 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
  2985
  (* 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
  2986
  (* below is more efficient                                             *)
16050
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15794
diff changeset
  2987
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  2988
  fun gfp_interpreter thy model args t =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2989
    case t of
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  2990
      Const (@{const_name gfp}, Type ("fun", [Type ("fun",
30312
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  2991
        [Type ("fun", [T, Type ("bool", [])]),
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  2992
         Type ("fun", [_, Type ("bool", [])])]),
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  2993
         Type ("fun", [_, Type ("bool", [])])])) =>
33037
b22e44496dc2 replaced old_style infixes eq_set, subset, union, inter and variants by generic versions
haftmann
parents: 33002
diff changeset
  2994
      let
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2995
        val size_elem = size_of_type thy model T
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2996
        (* the universe (i.e. the set that contains every element) *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  2997
        val i_univ = Node (replicate size_elem TT)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  2998
        (* all sets with elements from type 'T' *)
30312
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  2999
        val i_sets =
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  3000
          make_constants thy model (Type ("fun", [T, Type ("bool", [])]))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3001
        (* all functions that map sets to sets *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3002
        val i_funs = make_constants thy model (Type ("fun",
30312
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  3003
          [Type ("fun", [T, Type ("bool", [])]),
0e0cb7ac0281 Fix remaining occurrences of "'a set" in Refute, by using "'a => bool" instead.
blanchet
parents: 30307
diff changeset
  3004
           Type ("fun", [T, Type ("bool", [])])]))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3005
        (* "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
  3006
        (* interpretation * interpretation -> bool *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3007
        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
  3008
          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
  3009
            (subs ~~ sups)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3010
          | is_subset (_, _) =
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  3011
          raise REFUTE ("gfp_interpreter",
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3012
            "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
  3013
        (* interpretation * interpretation -> interpretation *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3014
        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
  3015
            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
  3016
                 (xs ~~ ys))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3017
          | union (_, _) =
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  3018
          raise REFUTE ("gfp_interpreter",
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3019
            "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
  3020
        (* interpretation -> interpretaion *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3021
        fun gfp (Node resultsets) =
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
  3022
          fold (fn (set, resultset) => fn acc =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3023
            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
  3024
              union (acc, set)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3025
            else
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
  3026
              acc) (i_sets ~~ resultsets) i_univ
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3027
          | gfp _ =
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  3028
            raise REFUTE ("gfp_interpreter",
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3029
              "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
  3030
      in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3031
        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
  3032
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3033
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3034
      NONE;
36130
9a672f7d488d commented out unsound "lfp"/"gfp" handling + fixed set output syntax;
blanchet
parents: 35746
diff changeset
  3035
*)
16050
828fc32f390f interpreters for lfp/gfp added
webertj
parents: 15794
diff changeset
  3036
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3037
  (* 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
  3038
    (interpretation * model * arguments) option *)
21267
5294ecae6708 interpreters for fst and snd added
webertj
parents: 21098
diff changeset
  3039
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3040
  (* 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
  3041
  (* 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
  3042
  (* below is more efficient                                             *)
21267
5294ecae6708 interpreters for fst and snd added
webertj
parents: 21098
diff changeset
  3043
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3044
  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
  3045
    case t of
29820
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
  3046
      Const (@{const_name fst}, Type ("fun", [Type ("*", [T, U]), _])) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3047
      let
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3048
        val constants_T = make_constants thy model T
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3049
        val size_U      = size_of_type thy model U
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3050
      in
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32950
diff changeset
  3051
        SOME (Node (maps (replicate size_U) constants_T), model, args)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3052
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3053
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3054
      NONE;
21267
5294ecae6708 interpreters for fst and snd added
webertj
parents: 21098
diff changeset
  3055
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3056
  (* 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
  3057
    (interpretation * model * arguments) option *)
21267
5294ecae6708 interpreters for fst and snd added
webertj
parents: 21098
diff changeset
  3058
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3059
  (* 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
  3060
  (* 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
  3061
  (* below is more efficient                                             *)
21267
5294ecae6708 interpreters for fst and snd added
webertj
parents: 21098
diff changeset
  3062
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3063
  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
  3064
    case t of
29820
07f53494cf20 Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
blanchet
parents: 29802
diff changeset
  3065
      Const (@{const_name snd}, Type ("fun", [Type ("*", [T, U]), _])) =>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3066
      let
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3067
        val size_T      = size_of_type thy model T
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3068
        val constants_U = make_constants thy model U
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3069
      in
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32950
diff changeset
  3070
        SOME (Node (flat (replicate size_T constants_U)), model, args)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3071
      end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3072
    | _ =>
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3073
      NONE;
21267
5294ecae6708 interpreters for fst and snd added
webertj
parents: 21098
diff changeset
  3074
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
  3075
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
  3076
(* ------------------------------------------------------------------------- *)
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
  3077
(* 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
  3078
(* ------------------------------------------------------------------------- *)
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
  3079
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3080
  (* theory -> model -> Term.typ -> interpretation -> (int -> bool) ->
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3081
    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
  3082
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3083
  fun stlc_printer thy model T intr assignment =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3084
  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3085
    (* string -> string *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3086
    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
  3087
      (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
  3088
        o explode) s
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3089
    (* Term.typ -> string *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3090
    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
  3091
      | 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
  3092
      | 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
  3093
      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
  3094
    (* interpretation -> int *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3095
    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
  3096
      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
  3097
      | index_from_interpretation _ =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3098
      raise REFUTE ("stlc_printer",
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3099
        "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
  3100
  in
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3101
    case T of
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3102
      Type ("fun", [T1, T2]) =>
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3103
      let
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3104
        (* create all constants of type 'T1' *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3105
        val constants = make_constants thy model T1
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3106
        (* interpretation list *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3107
        val results = (case intr of
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3108
            Node xs => xs
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3109
          | _       => raise REFUTE ("stlc_printer",
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3110
            "interpretation for function type is a leaf"))
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3111
        (* Term.term list *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3112
        val pairs = map (fn (arg, result) =>
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3113
          HOLogic.mk_prod
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3114
            (print thy model T1 arg assignment,
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3115
             print thy model T2 result assignment))
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3116
          (constants ~~ results)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3117
        (* Term.typ *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3118
        val HOLogic_prodT = HOLogic.mk_prodT (T1, T2)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3119
        val HOLogic_setT  = HOLogic.mk_setT HOLogic_prodT
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3120
        (* Term.term *)
36130
9a672f7d488d commented out unsound "lfp"/"gfp" handling + fixed set output syntax;
blanchet
parents: 35746
diff changeset
  3121
        val HOLogic_empty_set = Const (@{const_abbrev Set.empty}, HOLogic_setT)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3122
        val HOLogic_insert    =
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  3123
          Const (@{const_name insert}, HOLogic_prodT --> HOLogic_setT --> HOLogic_setT)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3124
      in
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
  3125
        SOME (fold_rev (fn pair => fn acc => HOLogic_insert $ pair $ acc) pairs HOLogic_empty_set)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3126
      end
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3127
    | Type ("prop", [])      =>
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3128
      (case index_from_interpretation intr of
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 26957
diff changeset
  3129
        ~1 => SOME (HOLogic.mk_Trueprop (Const (@{const_name undefined}, HOLogic.boolT)))
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3130
      | 0  => SOME (HOLogic.mk_Trueprop HOLogic.true_const)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3131
      | 1  => SOME (HOLogic.mk_Trueprop HOLogic.false_const)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3132
      | _  => raise REFUTE ("stlc_interpreter",
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3133
        "illegal interpretation for a propositional value"))
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3134
    | Type _  => if index_from_interpretation intr = (~1) then
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 26957
diff changeset
  3135
        SOME (Const (@{const_name undefined}, T))
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3136
      else
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3137
        SOME (Const (string_of_typ T ^
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3138
          string_of_int (index_from_interpretation intr), T))
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3139
    | TFree _ => if index_from_interpretation intr = (~1) then
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 26957
diff changeset
  3140
        SOME (Const (@{const_name undefined}, T))
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3141
      else
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3142
        SOME (Const (string_of_typ T ^
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3143
          string_of_int (index_from_interpretation intr), T))
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3144
    | TVar _  => if index_from_interpretation intr = (~1) then
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 26957
diff changeset
  3145
        SOME (Const (@{const_name undefined}, T))
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3146
      else
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3147
        SOME (Const (string_of_typ T ^
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3148
          string_of_int (index_from_interpretation intr), T))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3149
  end;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3150
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3151
  (* theory -> model -> Term.typ -> interpretation -> (int -> bool) ->
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3152
    Term.term option *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3153
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3154
  fun IDT_printer thy model T intr assignment =
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3155
    (case T of
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3156
      Type (s, Ts) =>
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31737
diff changeset
  3157
      (case Datatype.get_info thy s of
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3158
        SOME info =>  (* inductive datatype *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3159
        let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3160
          val (typs, _)           = model
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3161
          val index               = #index info
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3162
          val descr               = #descr info
29288
253bcf2a5854 avoid polymorphic equality;
wenzelm
parents: 29272
diff changeset
  3163
          val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3164
          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
  3165
          (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3166
          val _ = if Library.exists (fn d =>
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33955
diff changeset
  3167
              case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3168
            then
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3169
              raise REFUTE ("IDT_printer", "datatype argument (for type " ^
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26931
diff changeset
  3170
                Syntax.string_of_typ_global thy (Type (s, Ts)) ^ ") is not a variable")
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3171
            else ()
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3172
          (* 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
  3173
          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
  3174
              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
  3175
            | 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
  3176
              "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
  3177
        in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3178
          if element < 0 then
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 26957
diff changeset
  3179
            SOME (Const (@{const_name undefined}, Type (s, Ts)))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3180
          else let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3181
            (* 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
  3182
            (* 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
  3183
            (* '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
  3184
            (* indices of the arguments                                      *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3185
            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
  3186
              let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3187
                val cTerm      = Const (cname,
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3188
                  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
  3189
                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
  3190
                  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
  3191
                (* interpretation -> int list option *)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3192
                fun get_args (Leaf xs) =
31986
a68f88d264f7 dropped find_index_eq
haftmann
parents: 31943
diff changeset
  3193
                  if find_index (fn x => x = True) xs = element then
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3194
                    SOME []
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3195
                  else
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3196
                    NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3197
                  | get_args (Node xs) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3198
                  let
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3199
                    (* 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
  3200
                    fun search ([], _) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3201
                      NONE
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3202
                      | search (x::xs, n) =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3203
                      (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
  3204
                        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
  3205
                      | 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
  3206
                  in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3207
                    search (xs, 0)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3208
                  end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3209
              in
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3210
                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
  3211
              end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3212
            val (cTerm, cargs, args) =
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3213
              (* we could speed things up by computing the correct          *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3214
              (* constructor directly (rather than testing all              *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3215
              (* constructors), based on the order in which constructors    *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3216
              (* generate elements of datatypes; the current implementation *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3217
              (* of 'IDT_printer' however is independent of the internals   *)
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3218
              (* of 'IDT_constructor_interpreter'                           *)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3219
              (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
  3220
                SOME x => x
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3221
              | 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
  3222
                "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
  3223
                string_of_int element))
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3224
            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
  3225
              let
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3226
                val dT     = typ_of_dtyp descr typ_assoc d
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3227
                (* 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
  3228
                (* 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
  3229
                (* generate all constants                                 *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3230
                val consts = make_constants thy (typs, []) dT
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3231
              in
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3232
                print thy (typs, []) dT (List.nth (consts, n)) assignment
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3233
              end) (cargs ~~ args)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3234
          in
33246
46e47aa1558f eliminated some old folds;
wenzelm
parents: 33243
diff changeset
  3235
            SOME (list_comb (cTerm, argsTerms))
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3236
          end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3237
        end
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3238
      | 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
  3239
        NONE)
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3240
    | _ =>  (* a (free or schematic) type variable *)
25014
b711b0af178e significant code overhaul, bugfix for inductive data types
webertj
parents: 24928
diff changeset
  3241
      NONE);
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3242
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3243
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3244
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  3245
(* 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
  3246
(* structure                                                                 *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3247
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3248
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3249
(* ------------------------------------------------------------------------- *)
14456
cca28ec5f9a6 support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
webertj
parents: 14351
diff changeset
  3250
(* 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
  3251
(*       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
  3252
(*       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
  3253
(*       subterms that are then passed to other interpreters!                *)
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3254
(* ------------------------------------------------------------------------- *)
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3255
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3256
  val setup =
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3257
     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
  3258
     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
  3259
     add_interpreter "HOLogic" HOLogic_interpreter #>
29956
62f931b257b7 Reintroduce set_interpreter for Collect and op :.
blanchet
parents: 29820
diff changeset
  3260
     add_interpreter "set"     set_interpreter #>
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3261
     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
  3262
     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
  3263
     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
  3264
     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
  3265
     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
  3266
     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
  3267
     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
  3268
     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
  3269
     add_interpreter "Nat_HOL.times"      Nat_times_interpreter #>
25032
f7095d7cb9a3 interpreter for List.append added again
webertj
parents: 25014
diff changeset
  3270
     add_interpreter "List.append" List_append_interpreter #>
36130
9a672f7d488d commented out unsound "lfp"/"gfp" handling + fixed set output syntax;
blanchet
parents: 35746
diff changeset
  3271
(* UNSOUND
29802
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  3272
     add_interpreter "lfp" lfp_interpreter #>
d201a838d0f7 Make some Refute functions public so I can use them in Nitpick,
blanchet
parents: 29288
diff changeset
  3273
     add_interpreter "gfp" gfp_interpreter #>
36130
9a672f7d488d commented out unsound "lfp"/"gfp" handling + fixed set output syntax;
blanchet
parents: 35746
diff changeset
  3274
*)
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 22093
diff changeset
  3275
     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
  3276
     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
  3277
     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
  3278
     add_printer "IDT"  IDT_printer;
14350
41b32020d0b3 Adding 'refute' to HOL.
webertj
parents:
diff changeset
  3279
22092
ab3dfcef6489 reformatted to 80 chars/line
webertj
parents: 22055
diff changeset
  3280
end  (* structure Refute *)