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