| author | berghofe | 
| Thu, 10 Jan 2008 19:10:08 +0100 | |
| changeset 25886 | 7753e0d81b7a | 
| parent 25538 | 58e8ba3b792b | 
| child 26328 | b2d6f520172c | 
| permissions | -rw-r--r-- | 
| 14350 | 1 | (* Title: HOL/Tools/refute.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Tjark Weber | |
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 4 | Copyright 2003-2007 | 
| 14350 | 5 | |
| 14965 | 6 | Finite model generation for HOL formulas, using a SAT solver. | 
| 14350 | 7 | *) | 
| 8 | ||
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 9 | (* ------------------------------------------------------------------------- *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 10 | (* Declares the 'REFUTE' signature as well as a structure 'Refute'. *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 11 | (* Documentation is available in the Isabelle/Isar theory 'HOL/Refute.thy'. *) | 
| 14350 | 12 | (* ------------------------------------------------------------------------- *) | 
| 13 | ||
| 14 | signature REFUTE = | |
| 15 | sig | |
| 16 | ||
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 17 | exception REFUTE of string * string | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 18 | |
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 19 | (* ------------------------------------------------------------------------- *) | 
| 14807 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 20 | (* Model/interpretation related code (translation HOL -> propositional logic *) | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 21 | (* ------------------------------------------------------------------------- *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 22 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 23 | type params | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 24 | type interpretation | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 25 | type model | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 26 | type arguments | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 27 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 28 | exception MAXVARS_EXCEEDED | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 29 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 30 | val add_interpreter : string -> (theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 31 | (interpretation * model * arguments) option) -> theory -> theory | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 32 | val add_printer : string -> (theory -> model -> Term.typ -> | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 33 | interpretation -> (int -> bool) -> Term.term option) -> theory -> theory | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 34 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 35 | val interpret : theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 36 | (interpretation * model * arguments) | 
| 14807 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 37 | |
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 38 | val print : theory -> model -> Term.typ -> interpretation -> | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 39 | (int -> bool) -> Term.term | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 40 | val print_model : theory -> model -> (int -> bool) -> string | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 41 | |
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 42 | (* ------------------------------------------------------------------------- *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 43 | (* Interface *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 44 | (* ------------------------------------------------------------------------- *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 45 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 46 | val set_default_param : (string * string) -> theory -> theory | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 47 | val get_default_param : theory -> string -> string option | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 48 | val get_default_params : theory -> (string * string) list | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 49 | val actual_params : theory -> (string * string) list -> params | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 50 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 51 | val find_model : theory -> params -> Term.term -> bool -> unit | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 52 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 53 | (* tries to find a model for a formula: *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 54 | val satisfy_term : theory -> (string * string) list -> Term.term -> unit | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 55 | (* tries to find a model that refutes a formula: *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 56 | val refute_term : theory -> (string * string) list -> Term.term -> unit | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 57 | val refute_subgoal : | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 58 | theory -> (string * string) list -> Thm.thm -> int -> unit | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 59 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 60 | val setup : theory -> theory | 
| 22092 | 61 | |
| 62 | end; (* signature REFUTE *) | |
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 63 | |
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 64 | structure Refute : REFUTE = | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 65 | struct | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 66 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 67 | open PropLogic; | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 68 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 69 | (* We use 'REFUTE' only for internal error conditions that should *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 70 | (* never occur in the first place (i.e. errors caused by bugs in our *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 71 | (* code). Otherwise (e.g. to indicate invalid input data) we use *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 72 | (* 'error'. *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 73 |   exception REFUTE of string * string;  (* ("in function", "cause") *)
 | 
| 14350 | 74 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 75 | (* should be raised by an interpreter when more variables would be *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 76 | (* required than allowed by 'maxvars' *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 77 | exception MAXVARS_EXCEEDED; | 
| 14350 | 78 | |
| 79 | (* ------------------------------------------------------------------------- *) | |
| 80 | (* TREES *) | |
| 81 | (* ------------------------------------------------------------------------- *) | |
| 82 | ||
| 83 | (* ------------------------------------------------------------------------- *) | |
| 84 | (* tree: implements an arbitrarily (but finitely) branching tree as a list *) | |
| 85 | (* of (lists of ...) elements *) | |
| 86 | (* ------------------------------------------------------------------------- *) | |
| 87 | ||
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 88 | datatype 'a tree = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 89 | Leaf of 'a | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 90 |     | Node of ('a tree) list;
 | 
| 14350 | 91 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 92 |   (* ('a -> 'b) -> 'a tree -> 'b tree *)
 | 
| 14350 | 93 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 94 | fun tree_map f tr = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 95 | case tr of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 96 | Leaf x => Leaf (f x) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 97 | | Node xs => Node (map (tree_map f) xs); | 
| 14350 | 98 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 99 |   (* ('a * 'b -> 'a) -> 'a * ('b tree) -> 'a *)
 | 
| 14350 | 100 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 101 | fun tree_foldl f = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 102 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 103 | fun itl (e, Leaf x) = f(e,x) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 104 | | itl (e, Node xs) = Library.foldl (tree_foldl f) (e,xs) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 105 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 106 | itl | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 107 | end; | 
| 14350 | 108 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 109 |   (* 'a tree * 'b tree -> ('a * 'b) tree *)
 | 
| 14350 | 110 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 111 | fun tree_pair (t1, t2) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 112 | case t1 of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 113 | Leaf x => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 114 | (case t2 of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 115 | Leaf y => Leaf (x,y) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 116 |         | Node _ => raise REFUTE ("tree_pair",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 117 | "trees are of different height (second tree is higher)")) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 118 | | Node xs => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 119 | (case t2 of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 120 | (* '~~' will raise an exception if the number of branches in *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 121 | (* both trees is different at the current node *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 122 | Node ys => Node (map tree_pair (xs ~~ ys)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 123 |         | Leaf _  => raise REFUTE ("tree_pair",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 124 | "trees are of different height (first tree is higher)")); | 
| 14350 | 125 | |
| 126 | (* ------------------------------------------------------------------------- *) | |
| 14807 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 127 | (* params: parameters that control the translation into a propositional *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 128 | (* formula/model generation *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 129 | (* *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 130 | (* The following parameters are supported (and required (!), except for *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 131 | (* "sizes"): *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 132 | (* *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 133 | (* Name Type Description *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 134 | (* *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 135 | (* "sizes" (string * int) list *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 136 | (* Size of ground types (e.g. 'a=2), or depth of IDTs. *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 137 | (* "minsize" int If >0, minimal size of each ground type/IDT depth. *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 138 | (* "maxsize" int If >0, maximal size of each ground type/IDT depth. *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 139 | (* "maxvars" int If >0, use at most 'maxvars' Boolean variables *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 140 | (* when transforming the term into a propositional *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 141 | (* formula. *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 142 | (* "maxtime" int If >0, terminate after at most 'maxtime' seconds. *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 143 | (* "satsolver" string SAT solver to be used. *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 144 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 145 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 146 | type params = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 147 |     {
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 148 | sizes : (string * int) list, | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 149 | minsize : int, | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 150 | maxsize : int, | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 151 | maxvars : int, | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 152 | maxtime : int, | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 153 | satsolver: string | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
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: 
14604diff
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: 
14604diff
changeset | 156 | (* ------------------------------------------------------------------------- *) | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
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: 
14351diff
changeset | 158 | (* 'interpretation' *) | 
| 14350 | 159 | (* ------------------------------------------------------------------------- *) | 
| 160 | ||
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 161 | type interpretation = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 162 | prop_formula list tree; | 
| 14350 | 163 | |
| 164 | (* ------------------------------------------------------------------------- *) | |
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
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: 
14351diff
changeset | 166 | (* terms *) | 
| 14350 | 167 | (* ------------------------------------------------------------------------- *) | 
| 168 | ||
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 169 | type model = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 170 | (Term.typ * int) list * (Term.term * interpretation) list; | 
| 14350 | 171 | |
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 172 | (* ------------------------------------------------------------------------- *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 173 | (* arguments: additional arguments required during interpretation of terms *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
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: 
14604diff
changeset | 175 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 176 | type arguments = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 177 |     {
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 178 | (* just passed unchanged from 'params': *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 179 | maxvars : int, | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 180 | (* whether to use 'make_equality' or 'make_def_equality': *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 181 | def_eq : bool, | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 182 | (* the following may change during the translation: *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 183 | next_idx : int, | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 184 | bounds : interpretation list, | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 185 | wellformed: prop_formula | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 186 | }; | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 187 | |
| 14350 | 188 | |
| 22846 | 189 | structure RefuteData = TheoryDataFun | 
| 190 | ( | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 191 | type T = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 192 |       {interpreters: (string * (theory -> model -> arguments -> Term.term ->
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 193 | (interpretation * model * arguments) option)) list, | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 194 | printers: (string * (theory -> model -> Term.typ -> interpretation -> | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 195 | (int -> bool) -> Term.term option)) list, | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 196 | parameters: string Symtab.table}; | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 197 |     val empty = {interpreters = [], printers = [], parameters = Symtab.empty};
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 198 | val copy = I; | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 199 | val extend = I; | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 200 | fun merge _ | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 201 |       ({interpreters = in1, printers = pr1, parameters = pa1},
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 202 |        {interpreters = in2, printers = pr2, parameters = pa2}) =
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 203 |       {interpreters = AList.merge (op =) (K true) (in1, in2),
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 204 | printers = AList.merge (op =) (K true) (pr1, pr2), | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 205 | parameters = Symtab.merge (op=) (pa1, pa2)}; | 
| 22846 | 206 | ); | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 207 | |
| 14350 | 208 | |
| 209 | (* ------------------------------------------------------------------------- *) | |
| 15334 
d5a92997dc1b
exception CANNOT_INTERPRET removed (not needed anymore since the stlc_interpreter can interpret any term)
 webertj parents: 
15333diff
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: 
15333diff
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: 
15333diff
changeset | 212 | (* track of the interpretation of subterms *) | 
| 14350 | 213 | (* ------------------------------------------------------------------------- *) | 
| 214 | ||
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 215 | (* theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 216 | (interpretation * model * arguments) *) | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 217 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 218 | fun interpret thy model args t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 219 | case get_first (fn (_, f) => f thy model args t) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 220 | (#interpreters (RefuteData.get thy)) of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 221 |       NONE   => raise REFUTE ("interpret",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 222 | "no interpreter for term " ^ quote (Sign.string_of_term thy t)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 223 | | SOME x => x; | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 224 | |
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 225 | (* ------------------------------------------------------------------------- *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 226 | (* print: converts the interpretation 'intr', which must denote a term of *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 227 | (* type 'T', into a term using a suitable printer *) | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 228 | (* ------------------------------------------------------------------------- *) | 
| 14350 | 229 | |
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 230 | (* theory -> model -> Term.typ -> interpretation -> (int -> bool) -> | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 231 | Term.term *) | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 232 | |
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 233 | fun print thy model T intr assignment = | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 234 | case get_first (fn (_, f) => f thy model T intr assignment) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 235 | (#printers (RefuteData.get thy)) of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 236 |       NONE   => raise REFUTE ("print",
 | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 237 | "no printer for type " ^ quote (Sign.string_of_typ thy T)) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 238 | | SOME x => x; | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 239 | |
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 240 | (* ------------------------------------------------------------------------- *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 241 | (* 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: 
14604diff
changeset | 242 | (* (given by an assignment for Boolean variables) and suitable *) | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 243 | (* printers *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 244 | (* ------------------------------------------------------------------------- *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 245 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 246 | (* theory -> model -> (int -> bool) -> string *) | 
| 14807 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 247 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 248 | fun print_model thy model assignment = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 249 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 250 | val (typs, terms) = model | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 251 | val typs_msg = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 252 | if null typs then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 253 | "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: 
22093diff
changeset | 254 | else | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 255 | "Size of types: " ^ commas (map (fn (T, i) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 256 | Sign.string_of_typ thy T ^ ": " ^ string_of_int i) typs) ^ "\n" | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 257 | val show_consts_msg = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 258 | if not (!show_consts) andalso Library.exists (is_Const o fst) terms then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 259 | "set \"show_consts\" to show the interpretation of constants\n" | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 260 | else | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 261 | "" | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 262 | val terms_msg = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 263 | if null terms then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 264 | "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: 
22093diff
changeset | 265 | else | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 266 | space_implode "\n" (List.mapPartial (fn (t, intr) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 267 | (* print constants only if 'show_consts' is true *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 268 | if (!show_consts) orelse not (is_Const t) then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 269 | SOME (Sign.string_of_term thy t ^ ": " ^ | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 270 | Sign.string_of_term thy | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 271 | (print thy model (Term.type_of t) intr assignment)) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 272 | else | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 273 | NONE) terms) ^ "\n" | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 274 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 275 | typs_msg ^ show_consts_msg ^ terms_msg | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 276 | end; | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 277 | |
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 278 | |
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 279 | (* ------------------------------------------------------------------------- *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 280 | (* PARAMETER MANAGEMENT *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 281 | (* ------------------------------------------------------------------------- *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 282 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 283 | (* string -> (theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 284 | (interpretation * model * arguments) option) -> theory -> theory *) | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 285 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 286 | fun add_interpreter name f thy = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 287 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 288 |     val {interpreters, printers, parameters} = RefuteData.get thy
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 289 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 290 | case AList.lookup (op =) interpreters name of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 291 |       NONE   => RefuteData.put {interpreters = (name, f) :: interpreters,
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 292 | printers = printers, parameters = parameters} thy | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 293 |     | SOME _ => error ("Interpreter " ^ name ^ " already declared")
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 294 | end; | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 295 | |
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 296 | (* string -> (theory -> model -> Term.typ -> interpretation -> | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 297 | (int -> bool) -> Term.term option) -> theory -> theory *) | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 298 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 299 | fun add_printer name f thy = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 300 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 301 |     val {interpreters, printers, parameters} = RefuteData.get thy
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 302 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 303 | case AList.lookup (op =) printers name of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 304 |       NONE   => RefuteData.put {interpreters = interpreters,
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 305 | printers = (name, f) :: printers, parameters = parameters} thy | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 306 |     | SOME _ => error ("Printer " ^ name ^ " already declared")
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 307 | end; | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 308 | |
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 309 | (* ------------------------------------------------------------------------- *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 310 | (* set_default_param: stores the '(name, value)' pair in RefuteData's *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 311 | (* parameter table *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 312 | (* ------------------------------------------------------------------------- *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 313 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 314 | (* (string * string) -> theory -> theory *) | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 315 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 316 | fun set_default_param (name, value) thy = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 317 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 318 |     val {interpreters, printers, parameters} = RefuteData.get thy
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 319 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 320 | RefuteData.put (case Symtab.lookup parameters name of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 321 | NONE => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 322 |       {interpreters = interpreters, printers = printers,
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 323 | parameters = Symtab.extend (parameters, [(name, value)])} | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 324 | | SOME _ => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 325 |       {interpreters = interpreters, printers = printers,
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 326 | parameters = Symtab.update (name, value) parameters}) thy | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 327 | end; | 
| 14350 | 328 | |
| 329 | (* ------------------------------------------------------------------------- *) | |
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 330 | (* get_default_param: retrieves the value associated with 'name' from *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 331 | (* RefuteData's parameter table *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 332 | (* ------------------------------------------------------------------------- *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 333 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 334 | (* theory -> string -> string option *) | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 335 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 336 | val get_default_param = Symtab.lookup o #parameters o RefuteData.get; | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 337 | |
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 338 | (* ------------------------------------------------------------------------- *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 339 | (* get_default_params: returns a list of all '(name, value)' pairs that are *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 340 | (* stored in RefuteData's parameter table *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 341 | (* ------------------------------------------------------------------------- *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 342 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 343 | (* theory -> (string * string) list *) | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 344 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 345 | val get_default_params = Symtab.dest o #parameters o RefuteData.get; | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 346 | |
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 347 | (* ------------------------------------------------------------------------- *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 348 | (* actual_params: takes a (possibly empty) list 'params' of parameters that *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 349 | (* override the default parameters currently specified in 'thy', and *) | 
| 14807 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 350 | (* returns a record that can be passed to 'find_model'. *) | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 351 | (* ------------------------------------------------------------------------- *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 352 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 353 | (* theory -> (string * string) list -> params *) | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 354 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 355 | fun actual_params thy override = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 356 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 357 | (* (string * string) list * string -> int *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 358 | fun read_int (parms, name) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 359 | case AList.lookup (op =) parms name of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 360 | SOME s => (case Int.fromString s of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 361 | SOME i => i | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 362 |         | NONE   => error ("parameter " ^ quote name ^
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 363 | " (value is " ^ quote s ^ ") must be an integer value")) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 364 |       | NONE   => error ("parameter " ^ quote name ^
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 365 | " must be assigned a value") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 366 | (* (string * string) list * string -> string *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 367 | fun read_string (parms, name) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 368 | case AList.lookup (op =) parms name of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 369 | SOME s => s | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 370 |       | NONE   => error ("parameter " ^ quote name ^
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 371 | " must be assigned a value") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 372 | (* 'override' first, defaults last: *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 373 | (* (string * string) list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 374 | val allparams = override @ (get_default_params thy) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 375 | (* int *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 376 | val minsize = read_int (allparams, "minsize") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 377 | val maxsize = read_int (allparams, "maxsize") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 378 | val maxvars = read_int (allparams, "maxvars") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 379 | val maxtime = read_int (allparams, "maxtime") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 380 | (* string *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 381 | val satsolver = read_string (allparams, "satsolver") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 382 | (* 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: 
22093diff
changeset | 383 | (* 'sizes' *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 384 | (* 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: 
22093diff
changeset | 385 | (* 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: 
22093diff
changeset | 386 | (* (string * int) list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 387 | val sizes = List.mapPartial | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 388 | (fn (name, value) => Option.map (pair name) (Int.fromString value)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 389 | (List.filter (fn (name, _) => name<>"minsize" andalso name<>"maxsize" | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 390 | andalso name<>"maxvars" andalso name<>"maxtime" | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 391 | andalso name<>"satsolver") allparams) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 392 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 393 |     {sizes=sizes, minsize=minsize, maxsize=maxsize, maxvars=maxvars,
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 394 | maxtime=maxtime, satsolver=satsolver} | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 395 | 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: 
14604diff
changeset | 396 | |
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 397 | |
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 398 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 399 | (* 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: 
14604diff
changeset | 400 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 401 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 402 |   (* (''a * 'b) list -> ''a -> 'b *)
 | 
| 22092 | 403 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 404 | fun lookup xs key = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 405 | Option.valOf (AList.lookup (op =) xs key); | 
| 22092 | 406 | |
| 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: 
14604diff
changeset | 407 | (* ------------------------------------------------------------------------- *) | 
| 15335 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 408 | (* typ_of_dtyp: converts a data type ('DatatypeAux.dtyp') into a type        *)
 | 
| 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 409 | (*              ('Term.typ'), given type parameters for the data type's type *)
 | 
| 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 410 | (* arguments *) | 
| 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 411 | (* ------------------------------------------------------------------------- *) | 
| 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 412 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 413 | (* DatatypeAux.descr -> (DatatypeAux.dtyp * Term.typ) list -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 414 | DatatypeAux.dtyp -> Term.typ *) | 
| 15335 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 415 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 416 | fun typ_of_dtyp descr typ_assoc (DatatypeAux.DtTFree a) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 417 | (* replace a 'DtTFree' variable by the associated type *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 418 | lookup typ_assoc (DatatypeAux.DtTFree a) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 419 | | typ_of_dtyp descr typ_assoc (DatatypeAux.DtType (s, ds)) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 420 | Type (s, map (typ_of_dtyp descr typ_assoc) ds) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 421 | | typ_of_dtyp descr typ_assoc (DatatypeAux.DtRec i) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 422 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 423 | val (s, ds, _) = lookup descr i | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 424 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 425 | Type (s, map (typ_of_dtyp descr typ_assoc) ds) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 426 | end; | 
| 15335 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 427 | |
| 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 428 | (* ------------------------------------------------------------------------- *) | 
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 429 | (* close_form: universal closure over schematic variables in 't' *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 430 | (* ------------------------------------------------------------------------- *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 431 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 432 | (* Term.term -> Term.term *) | 
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 433 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 434 | fun close_form t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 435 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 436 | (* (Term.indexname * Term.typ) list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 437 | val vars = sort_wrt (fst o fst) (map dest_Var (term_vars t)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 438 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 439 | Library.foldl (fn (t', ((x, i), T)) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 440 | (Term.all T) $ Abs (x, T, abstract_over (Var ((x, i), T), t'))) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 441 | (t, vars) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 442 | end; | 
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 443 | |
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 444 | (* ------------------------------------------------------------------------- *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 445 | (* monomorphic_term: applies a type substitution 'typeSubs' for all type *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 446 | (* variables in a term 't' *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 447 | (* ------------------------------------------------------------------------- *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 448 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 449 | (* Type.tyenv -> Term.term -> Term.term *) | 
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 450 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 451 | fun monomorphic_term typeSubs t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 452 | map_types (map_type_tvar | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 453 | (fn v => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 454 | case Type.lookup (typeSubs, v) of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 455 | NONE => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 456 | (* schematic type variable not instantiated *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 457 |           raise REFUTE ("monomorphic_term",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 458 | "no substitution for type variable " ^ fst (fst v) ^ | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 459 | " in term " ^ Display.raw_string_of_term t) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 460 | | SOME typ => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 461 | typ)) t; | 
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 462 | |
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 463 | (* ------------------------------------------------------------------------- *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 464 | (* specialize_type: given a constant 's' of type 'T', which is a subterm of *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 465 | (* 't', where 't' has a (possibly) more general type, the *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 466 | (* schematic type variables in 't' are instantiated to *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 467 | (* match the type 'T' (may raise Type.TYPE_MATCH) *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 468 | (* ------------------------------------------------------------------------- *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 469 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 470 | (* theory -> (string * Term.typ) -> Term.term -> Term.term *) | 
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 471 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 472 | fun specialize_type thy (s, T) t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 473 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 474 | fun find_typeSubs (Const (s', T')) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 475 | if s=s' then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 476 | SOME (Sign.typ_match thy (T', T) Vartab.empty) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 477 | handle Type.TYPE_MATCH => NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 478 | else | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 479 | NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 480 | | find_typeSubs (Free _) = NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 481 | | find_typeSubs (Var _) = NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 482 | | find_typeSubs (Bound _) = NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 483 | | find_typeSubs (Abs (_, _, body)) = find_typeSubs body | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 484 | | find_typeSubs (t1 $ t2) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 485 | (case find_typeSubs t1 of SOME x => SOME x | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 486 | | NONE => find_typeSubs t2) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 487 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 488 | case find_typeSubs t of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 489 | SOME typeSubs => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 490 | monomorphic_term typeSubs t | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 491 | | NONE => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 492 | (* no match found - perhaps due to sort constraints *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 493 | raise Type.TYPE_MATCH | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 494 | end; | 
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 495 | |
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 496 | (* ------------------------------------------------------------------------- *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 497 | (* 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: 
21931diff
changeset | 498 | (* denotes membership to an axiomatic type class *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 499 | (* ------------------------------------------------------------------------- *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 500 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 501 | (* theory -> string * Term.typ -> bool *) | 
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 502 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 503 | fun is_const_of_class thy (s, T) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 504 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 505 | 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: 
22093diff
changeset | 506 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 507 | (* 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: 
22093diff
changeset | 508 | (* or if we should also check the type 'T'. *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 509 | s mem_string class_const_names | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 510 | end; | 
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 511 | |
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 512 | (* ------------------------------------------------------------------------- *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 513 | (* 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: 
21931diff
changeset | 514 | (* of an inductive datatype in 'thy' *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 515 | (* ------------------------------------------------------------------------- *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 516 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 517 | (* theory -> string * Term.typ -> bool *) | 
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 518 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 519 | fun is_IDT_constructor thy (s, T) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 520 | (case body_type T of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 521 | Type (s', _) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 522 | (case DatatypePackage.get_datatype_constrs thy s' of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 523 | SOME constrs => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 524 | List.exists (fn (cname, cty) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 525 | cname = s andalso Sign.typ_instance thy (T, cty)) constrs | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 526 | | NONE => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 527 | false) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 528 | | _ => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 529 | false); | 
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 530 | |
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 531 | (* ------------------------------------------------------------------------- *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 532 | (* 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: 
21931diff
changeset | 533 | (* operator of an inductive datatype in 'thy' *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 534 | (* ------------------------------------------------------------------------- *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 535 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 536 | (* theory -> string * Term.typ -> bool *) | 
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 537 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 538 | fun is_IDT_recursor thy (s, T) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 539 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 540 | val rec_names = Symtab.fold (append o #rec_names o snd) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 541 | (DatatypePackage.get_datatypes thy) [] | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 542 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 543 | (* 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: 
22093diff
changeset | 544 | (* or if we should also check the type 'T'. *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 545 | s mem_string rec_names | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 546 | end; | 
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 547 | |
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 548 | (* ------------------------------------------------------------------------- *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 549 | (* get_def: looks up the definition of a constant, as created by "constdefs" *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 550 | (* ------------------------------------------------------------------------- *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 551 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 552 | (* theory -> string * Term.typ -> (string * Term.term) option *) | 
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 553 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 554 | fun get_def thy (s, T) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 555 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 556 | (* maps f ?t1 ... ?tn == rhs to %t1...tn. rhs *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 557 | fun norm_rhs eqn = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 558 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 559 | fun lambda (v as Var ((x, _), T)) t = Abs (x, T, abstract_over (v, t)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 560 |         | lambda v t                      = raise TERM ("lambda", [v, t])
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 561 | val (lhs, rhs) = Logic.dest_equals eqn | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 562 | val (_, args) = Term.strip_comb lhs | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 563 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 564 | fold lambda (rev args) rhs | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 565 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 566 | (* (string * Term.term) list -> (string * Term.term) option *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 567 | fun get_def_ax [] = NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 568 | | get_def_ax ((axname, ax) :: axioms) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 569 | (let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 570 | val (lhs, _) = Logic.dest_equals ax (* equations only *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 571 | val c = Term.head_of lhs | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 572 | val (s', T') = Term.dest_Const c | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 573 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 574 | if s=s' then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 575 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 576 | val typeSubs = Sign.typ_match thy (T', T) Vartab.empty | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 577 | val ax' = monomorphic_term typeSubs ax | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 578 | val rhs = norm_rhs ax' | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 579 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 580 | SOME (axname, rhs) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 581 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 582 | else | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 583 | get_def_ax axioms | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 584 | end handle ERROR _ => get_def_ax axioms | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 585 | | TERM _ => get_def_ax axioms | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 586 | | Type.TYPE_MATCH => get_def_ax axioms) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 587 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 588 | 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: 
22093diff
changeset | 589 | end; | 
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 590 | |
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 591 | (* ------------------------------------------------------------------------- *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 592 | (* 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: 
21931diff
changeset | 593 | (* ------------------------------------------------------------------------- *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 594 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 595 | (* theory -> (string * Term.typ) -> (string * Term.term) option *) | 
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 596 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 597 | fun get_typedef thy T = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 598 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 599 | (* (string * Term.term) list -> (string * Term.term) option *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 600 | fun get_typedef_ax [] = NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 601 | | get_typedef_ax ((axname, ax) :: axioms) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 602 | (let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 603 | (* Term.term -> Term.typ option *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 604 | fun type_of_type_definition (Const (s', T')) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 605 | if s'="Typedef.type_definition" then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 606 | SOME T' | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 607 | else | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 608 | NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 609 | | type_of_type_definition (Free _) = NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 610 | | type_of_type_definition (Var _) = NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 611 | | type_of_type_definition (Bound _) = NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 612 | | type_of_type_definition (Abs (_, _, body)) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 613 | type_of_type_definition body | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 614 | | type_of_type_definition (t1 $ t2) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 615 | (case type_of_type_definition t1 of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 616 | SOME x => SOME x | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 617 | | NONE => type_of_type_definition t2) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 618 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 619 | case type_of_type_definition ax of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 620 | SOME T' => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 621 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 622 | val T'' = (domain_type o domain_type) T' | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 623 | val typeSubs = Sign.typ_match thy (T'', T) Vartab.empty | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 624 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 625 | SOME (axname, monomorphic_term typeSubs ax) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 626 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 627 | | NONE => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 628 | get_typedef_ax axioms | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 629 | end handle ERROR _ => get_typedef_ax axioms | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 630 | | MATCH => get_typedef_ax axioms | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 631 | | Type.TYPE_MATCH => get_typedef_ax axioms) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 632 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 633 | 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: 
22093diff
changeset | 634 | end; | 
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 635 | |
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 636 | (* ------------------------------------------------------------------------- *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 637 | (* 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: 
21931diff
changeset | 638 | (* created by the "axclass" command *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 639 | (* ------------------------------------------------------------------------- *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 640 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 641 | (* theory -> string -> (string * Term.term) option *) | 
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 642 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 643 | fun get_classdef thy class = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 644 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 645 | val axname = class ^ "_class_def" | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 646 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 647 | Option.map (pair axname) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 648 | (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: 
22093diff
changeset | 649 | end; | 
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 650 | |
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 651 | (* ------------------------------------------------------------------------- *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 652 | (* 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: 
21931diff
changeset | 653 | (* normalizes the result term; certain constants are not *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 654 | (* unfolded (cf. 'collect_axioms' and the various interpreters *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 655 | (* below): if the interpretation respects a definition anyway, *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 656 | (* that definition does not need to be unfolded *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 657 | (* ------------------------------------------------------------------------- *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 658 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 659 | (* theory -> Term.term -> Term.term *) | 
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 660 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 661 | (* Note: we could intertwine unfolding of constants and beta-(eta-) *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 662 | (* normalization; this would save some unfolding for terms where *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 663 | (* constants are eliminated by beta-reduction (e.g. 'K c1 c2'). On *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 664 | (* the other hand, this would cause additional work for terms where *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 665 | (* 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: 
21931diff
changeset | 666 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 667 | fun unfold_defs thy t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 668 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 669 | (* Term.term -> Term.term *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 670 | fun unfold_loop t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 671 | case t of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 672 | (* Pure *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 673 |         Const ("all", _)                => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 674 |       | Const ("==", _)                 => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 675 |       | Const ("==>", _)                => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 676 |       | Const ("TYPE", _)               => t  (* axiomatic type classes *)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 677 | (* HOL *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 678 |       | Const ("Trueprop", _)           => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 679 |       | Const ("Not", _)                => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 680 | | (* redundant, since 'True' is also an IDT constructor *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 681 |         Const ("True", _)               => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 682 | | (* redundant, since 'False' is also an IDT constructor *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 683 |         Const ("False", _)              => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 684 |       | Const ("arbitrary", _)          => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 685 |       | Const ("The", _)                => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 686 |       | Const ("Hilbert_Choice.Eps", _) => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 687 |       | Const ("All", _)                => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 688 |       | Const ("Ex", _)                 => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 689 |       | Const ("op =", _)               => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 690 |       | Const ("op &", _)               => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 691 |       | Const ("op |", _)               => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 692 |       | Const ("op -->", _)             => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 693 | (* sets *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 694 |       | Const ("Collect", _)            => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 695 |       | Const ("op :", _)               => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 696 | (* other optimizations *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 697 |       | Const ("Finite_Set.card", _)    => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 698 |       | Const ("Finite_Set.Finites", _) => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 699 |       | Const ("Finite_Set.finite", _)  => t
 | 
| 23881 | 700 |       | Const (@{const_name HOL.less}, Type ("fun", [Type ("nat", []),
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 701 |         Type ("fun", [Type ("nat", []), Type ("bool", [])])])) => t
 | 
| 22997 | 702 |       | Const (@{const_name HOL.plus}, Type ("fun", [Type ("nat", []),
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 703 |         Type ("fun", [Type ("nat", []), Type ("nat", [])])])) => t
 | 
| 22997 | 704 |       | Const (@{const_name HOL.minus}, Type ("fun", [Type ("nat", []),
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 705 |         Type ("fun", [Type ("nat", []), Type ("nat", [])])])) => t
 | 
| 22997 | 706 |       | Const (@{const_name HOL.times}, Type ("fun", [Type ("nat", []),
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 707 |         Type ("fun", [Type ("nat", []), Type ("nat", [])])])) => t
 | 
| 25032 | 708 |       | Const ("List.append", _)        => t
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 709 |       | Const ("Lfp.lfp", _)            => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 710 |       | Const ("Gfp.gfp", _)            => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 711 |       | Const ("fst", _)                => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 712 |       | Const ("snd", _)                => t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 713 | (* simply-typed lambda calculus *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 714 | | Const (s, T) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 715 | (if is_IDT_constructor thy (s, T) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 716 | orelse is_IDT_recursor thy (s, T) then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 717 | t (* do not unfold IDT constructors/recursors *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 718 | (* unfold the constant if there is a defining equation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 719 | else case get_def thy (s, T) of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 720 | SOME (axname, rhs) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 721 | (* Note: if the term to be unfolded (i.e. 'Const (s, T)') *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 722 | (* occurs on the right-hand side of the equation, i.e. in *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 723 | (* 'rhs', we must not use this equation to unfold, because *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 724 | (* that would loop. Here would be the right place to *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 725 | (* check this. However, getting this really right seems *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 726 | (* difficult because the user may state arbitrary axioms, *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 727 | (* which could interact with overloading to create loops. *) | 
| 22580 | 728 |           ((*Output.immediate_output (" unfolding: " ^ axname);*)unfold_loop rhs)
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 729 | | NONE => t) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 730 | | Free _ => t | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 731 | | Var _ => t | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 732 | | Bound _ => t | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 733 | | Abs (s, T, body) => Abs (s, T, unfold_loop body) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 734 | | t1 $ t2 => (unfold_loop t1) $ (unfold_loop t2) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 735 | 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: 
22093diff
changeset | 736 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 737 | result | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 738 | end; | 
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 739 | |
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 740 | (* ------------------------------------------------------------------------- *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 741 | (* collect_axioms: collects (monomorphic, universally quantified, unfolded *) | 
| 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 742 | (* 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: 
14604diff
changeset | 743 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 744 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 745 | (* Note: to make the collection of axioms more easily extensible, this *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 746 | (* function could be based on user-supplied "axiom collectors", *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 747 | (* 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: 
14604diff
changeset | 748 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 749 | (* Note: currently we use "inverse" functions to the definitional *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 750 | (* mechanisms provided by Isabelle/HOL, e.g. for "axclass", *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 751 | (* "typedef", "constdefs". A more general approach could consider *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 752 | (* *every* axiom of the theory and collect it if it has a constant/ *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 753 | (* type/typeclass in common with the term 't'. *) | 
| 21985 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 webertj parents: 
21931diff
changeset | 754 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 755 | (* theory -> Term.term -> Term.term list *) | 
| 14807 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 756 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 757 | (* Which axioms are "relevant" for a particular term/type goes hand in *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 758 | (* hand with the interpretation of that term/type by its interpreter (see *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 759 | (* way below): if the interpretation respects an axiom anyway, the axiom *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 760 | (* 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: 
14604diff
changeset | 761 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 762 | (* To avoid collecting the same axiom multiple times, we use an *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 763 | (* 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: 
14604diff
changeset | 764 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 765 | fun collect_axioms thy t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 766 | let | 
| 22580 | 767 | val _ = Output.immediate_output "Adding axioms..." | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 768 | (* (string * Term.term) list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 769 | val axioms = Theory.all_axioms_of thy | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 770 | (* string * Term.term -> Term.term list -> Term.term list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 771 | fun collect_this_axiom (axname, ax) axs = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 772 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 773 | val ax' = unfold_defs thy ax | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 774 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 775 | if member (op aconv) axs ax' then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 776 | axs | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 777 | else ( | 
| 22580 | 778 |         Output.immediate_output (" " ^ axname);
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 779 | collect_term_axioms (ax' :: axs, ax') | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 780 | ) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 781 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 782 | (* Term.term list * Term.typ -> Term.term list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 783 | and collect_sort_axioms (axs, T) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 784 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 785 | (* string list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 786 | val sort = (case T of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 787 | TFree (_, sort) => sort | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 788 | | TVar (_, sort) => sort | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 789 |         | _               => raise REFUTE ("collect_axioms", "type " ^
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 790 | Sign.string_of_typ thy T ^ " is not a variable")) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 791 | (* obtain axioms for all superclasses *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 792 | val superclasses = sort @ (maps (Sign.super_classes thy) sort) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 793 | (* merely an optimization, because 'collect_this_axiom' disallows *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 794 | (* duplicate axioms anyway: *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 795 | val superclasses = distinct (op =) superclasses | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 796 | val class_axioms = maps (fn class => map (fn ax => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 797 |         ("<" ^ class ^ ">", Thm.prop_of ax))
 | 
| 24928 
3419943838f5
renamed AxClass.get_definition to AxClass.get_info (again);
 wenzelm parents: 
24688diff
changeset | 798 | (#axioms (AxClass.get_info thy class) handle ERROR _ => [])) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 799 | superclasses | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 800 | (* replace the (at most one) schematic type variable in each axiom *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 801 | (* by the actual type 'T' *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 802 | val monomorphic_class_axioms = map (fn (axname, ax) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 803 | (case Term.term_tvars ax of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 804 | [] => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 805 | (axname, ax) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 806 | | [(idx, S)] => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 807 | (axname, monomorphic_term (Vartab.make [(idx, (S, T))]) ax) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 808 | | _ => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 809 |           raise REFUTE ("collect_axioms", "class axiom " ^ axname ^ " (" ^
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 810 | Sign.string_of_term thy ax ^ | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 811 | ") contains more than one type variable"))) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 812 | class_axioms | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 813 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 814 | fold collect_this_axiom monomorphic_class_axioms axs | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 815 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 816 | (* Term.term list * Term.typ -> Term.term list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 817 | and collect_type_axioms (axs, T) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 818 | case T of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 819 | (* simple types *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 820 |         Type ("prop", [])      => axs
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 821 |       | Type ("fun", [T1, T2]) => collect_type_axioms
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 822 | (collect_type_axioms (axs, T1), T2) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 823 |       | Type ("set", [T1])     => collect_type_axioms (axs, T1)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 824 | (* axiomatic type classes *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 825 |       | Type ("itself", [T1])  => collect_type_axioms (axs, T1)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 826 | | Type (s, Ts) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 827 | (case DatatypePackage.get_datatype thy s of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 828 | SOME info => (* inductive datatype *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 829 | (* only collect relevant type axioms for the argument types *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 830 | Library.foldl collect_type_axioms (axs, Ts) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 831 | | NONE => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 832 | (case get_typedef thy T of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 833 | SOME (axname, ax) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 834 | collect_this_axiom (axname, ax) axs | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 835 | | NONE => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 836 | (* unspecified type, perhaps introduced with "typedecl" *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 837 | (* at least collect relevant type axioms for the argument types *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 838 | Library.foldl collect_type_axioms (axs, Ts))) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 839 | (* axiomatic type classes *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 840 | | TFree _ => collect_sort_axioms (axs, T) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 841 | (* axiomatic type classes *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 842 | | TVar _ => collect_sort_axioms (axs, T) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 843 | (* Term.term list * Term.term -> Term.term list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 844 | and collect_term_axioms (axs, t) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 845 | case t of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 846 | (* Pure *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 847 |         Const ("all", _)                => axs
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 848 |       | Const ("==", _)                 => axs
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 849 |       | Const ("==>", _)                => axs
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 850 | (* axiomatic type classes *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 851 |       | Const ("TYPE", T)               => collect_type_axioms (axs, T)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 852 | (* HOL *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 853 |       | Const ("Trueprop", _)           => axs
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 854 |       | Const ("Not", _)                => axs
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 855 | (* redundant, since 'True' is also an IDT constructor *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 856 |       | Const ("True", _)               => axs
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 857 | (* redundant, since 'False' is also an IDT constructor *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 858 |       | Const ("False", _)              => axs
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 859 |       | Const ("arbitrary", T)          => collect_type_axioms (axs, T)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 860 |       | Const ("The", T)                =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 861 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 862 |           val ax = specialize_type thy ("The", T)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 863 | (lookup axioms "HOL.the_eq_trivial") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 864 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 865 |           collect_this_axiom ("HOL.the_eq_trivial", ax) axs
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 866 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 867 |       | Const ("Hilbert_Choice.Eps", T) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 868 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 869 |           val ax = specialize_type thy ("Hilbert_Choice.Eps", T)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 870 | (lookup axioms "Hilbert_Choice.someI") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 871 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 872 |           collect_this_axiom ("Hilbert_Choice.someI", ax) axs
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 873 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 874 |       | Const ("All", T)                => collect_type_axioms (axs, T)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 875 |       | Const ("Ex", T)                 => collect_type_axioms (axs, T)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 876 |       | Const ("op =", T)               => collect_type_axioms (axs, T)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 877 |       | Const ("op &", _)               => axs
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 878 |       | Const ("op |", _)               => axs
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 879 |       | Const ("op -->", _)             => axs
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 880 | (* sets *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 881 |       | Const ("Collect", T)            => collect_type_axioms (axs, T)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 882 |       | Const ("op :", T)               => collect_type_axioms (axs, T)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 883 | (* other optimizations *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 884 |       | Const ("Finite_Set.card", T)    => collect_type_axioms (axs, T)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 885 |       | Const ("Finite_Set.Finites", T) => collect_type_axioms (axs, T)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 886 |       | Const ("Finite_Set.finite", T)  => collect_type_axioms (axs, T)
 | 
| 23881 | 887 |       | Const (@{const_name HOL.less}, T as Type ("fun", [Type ("nat", []),
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 888 |         Type ("fun", [Type ("nat", []), Type ("bool", [])])])) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 889 | collect_type_axioms (axs, T) | 
| 22997 | 890 |       | Const (@{const_name HOL.plus}, T as Type ("fun", [Type ("nat", []),
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 891 |         Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 892 | collect_type_axioms (axs, T) | 
| 22997 | 893 |       | Const (@{const_name HOL.minus}, T as Type ("fun", [Type ("nat", []),
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 894 |         Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 895 | collect_type_axioms (axs, T) | 
| 22997 | 896 |       | Const (@{const_name HOL.times}, T as Type ("fun", [Type ("nat", []),
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 897 |         Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 898 | collect_type_axioms (axs, T) | 
| 25032 | 899 |       | Const ("List.append", T)        => collect_type_axioms (axs, T)
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 900 |       | Const ("Lfp.lfp", T)            => collect_type_axioms (axs, T)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 901 |       | Const ("Gfp.gfp", T)            => collect_type_axioms (axs, T)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 902 |       | Const ("fst", T)                => collect_type_axioms (axs, T)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 903 |       | Const ("snd", T)                => collect_type_axioms (axs, T)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 904 | (* simply-typed lambda calculus *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 905 | | Const (s, T) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 906 | 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: 
22093diff
changeset | 907 | (* 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: 
22093diff
changeset | 908 | (* and the class definition *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 909 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 910 | val class = Logic.class_of_const s | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 911 |               val inclass = Logic.mk_inclass (TVar (("'a", 0), [class]), class)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 912 | val ax_in = SOME (specialize_type thy (s, T) inclass) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 913 | (* 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: 
22093diff
changeset | 914 | handle Type.TYPE_MATCH => NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 915 | val ax_1 = Option.map (fn ax => (Sign.string_of_term thy ax, ax)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 916 | ax_in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 917 | val ax_2 = Option.map (apsnd (specialize_type thy (s, T))) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 918 | (get_classdef thy class) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 919 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 920 | collect_type_axioms (fold collect_this_axiom | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 921 | (map_filter I [ax_1, ax_2]) axs, T) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 922 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 923 | else if is_IDT_constructor thy (s, T) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 924 | orelse is_IDT_recursor thy (s, T) then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 925 | (* only collect relevant type axioms *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 926 | collect_type_axioms (axs, T) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 927 | else | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 928 | (* 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: 
22093diff
changeset | 929 | (* 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: 
22093diff
changeset | 930 | (* typedefs, or type-class related constants *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 931 | (* only collect relevant type axioms *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 932 | collect_type_axioms (axs, T) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 933 | | Free (_, T) => collect_type_axioms (axs, T) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 934 | | Var (_, T) => collect_type_axioms (axs, T) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 935 | | Bound i => axs | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 936 | | Abs (_, T, body) => collect_term_axioms | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 937 | (collect_type_axioms (axs, T), body) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 938 | | t1 $ t2 => collect_term_axioms | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 939 | (collect_term_axioms (axs, t1), t2) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 940 | (* Term.term list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 941 | val result = map close_form (collect_term_axioms ([], t)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 942 | val _ = writeln " ...done." | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 943 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 944 | result | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 945 | end; | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 946 | |
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 947 | (* ------------------------------------------------------------------------- *) | 
| 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: 
14604diff
changeset | 948 | (* 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: 
14604diff
changeset | 949 | (* 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: 
14604diff
changeset | 950 | (* 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: 
14604diff
changeset | 951 | (* 'propT'. For IDTs, also the argument types of constructors *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 952 | (* 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: 
14604diff
changeset | 953 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 954 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 955 | (* theory -> Term.term -> Term.typ list *) | 
| 14807 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 956 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 957 | fun ground_types thy t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 958 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 959 | (* Term.typ * Term.typ list -> Term.typ list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 960 | fun collect_types (T, acc) = | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 961 | (case T of | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 962 |         Type ("fun", [T1, T2]) => collect_types (T1, collect_types (T2, acc))
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 963 |       | Type ("prop", [])      => acc
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 964 |       | Type ("set", [T1])     => collect_types (T1, acc)
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 965 | | Type (s, Ts) => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 966 | (case DatatypePackage.get_datatype thy s of | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 967 | SOME info => (* inductive datatype *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 968 | let | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 969 | val index = #index info | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 970 | val descr = #descr info | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 971 | val (_, typs, _) = lookup descr index | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 972 | val typ_assoc = typs ~~ Ts | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 973 | (* sanity check: every element in 'dtyps' must be a *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 974 | (* 'DtTFree' *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 975 | val _ = if Library.exists (fn d => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 976 | case d of DatatypeAux.DtTFree _ => false | _ => true) typs then | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 977 |               raise REFUTE ("ground_types", "datatype argument (for type "
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 978 | ^ Sign.string_of_typ thy T ^ ") is not a variable") | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 979 | else () | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 980 | (* required for mutually recursive datatypes; those need to *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 981 | (* be added even if they are an instance of an otherwise non- *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 982 | (* recursive datatype *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 983 | fun collect_dtyp (d, acc) = | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 984 | let | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 985 | val dT = typ_of_dtyp descr typ_assoc d | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 986 | in | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 987 | case d of | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 988 | DatatypeAux.DtTFree _ => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 989 | collect_types (dT, acc) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 990 | | DatatypeAux.DtType (_, ds) => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 991 | collect_types (dT, foldr collect_dtyp acc ds) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 992 | | DatatypeAux.DtRec i => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 993 | if dT mem acc then | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 994 | acc (* prevent infinite recursion *) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 995 | else | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 996 | let | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 997 | val (_, dtyps, dconstrs) = lookup descr i | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 998 | (* if the current type is a recursive IDT (i.e. a depth *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 999 | (* is required), add it to 'acc' *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1000 | val acc_dT = if Library.exists (fn (_, ds) => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1001 | Library.exists DatatypeAux.is_rec_type ds) dconstrs then | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1002 | insert (op =) dT acc | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1003 | else acc | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1004 | (* collect argument types *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1005 | val acc_dtyps = foldr collect_dtyp acc_dT dtyps | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1006 | (* collect constructor types *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1007 | val acc_dconstrs = foldr collect_dtyp acc_dtyps | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1008 | (List.concat (map snd dconstrs)) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1009 | in | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1010 | acc_dconstrs | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1011 | end | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1012 | end | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1013 | in | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1014 | (* argument types 'Ts' could be added here, but they are also *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1015 | (* added by 'collect_dtyp' automatically *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1016 | collect_dtyp (DatatypeAux.DtRec index, acc) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1017 | end | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1018 | | NONE => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1019 | (* not an inductive datatype, e.g. defined via "typedef" or *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1020 | (* "typedecl" *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1021 | insert (op =) T (foldr collect_types acc Ts)) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1022 | | TFree _ => insert (op =) T acc | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1023 | | TVar _ => insert (op =) T acc) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1024 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1025 | it_term_types collect_types (t, []) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1026 | 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: 
14604diff
changeset | 1027 | |
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1028 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1029 | (* 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: 
14604diff
changeset | 1030 | (* 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: 
14604diff
changeset | 1031 | (* 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: 
14604diff
changeset | 1032 | (* 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: 
14604diff
changeset | 1033 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1034 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1035 | (* Term.typ -> string *) | 
| 14807 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1036 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1037 | fun string_of_typ (Type (s, _)) = s | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1038 | | string_of_typ (TFree (s, _)) = s | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1039 | | 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: 
14604diff
changeset | 1040 | |
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1041 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1042 | (* 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: 
14604diff
changeset | 1043 | (* '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: 
14604diff
changeset | 1044 | (* '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: 
14604diff
changeset | 1045 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1046 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1047 | (* Term.typ list -> (string * int) list -> int -> (Term.typ * int) list *) | 
| 14807 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1048 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1049 | fun first_universe xs sizes minsize = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1050 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1051 | fun size_of_typ T = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1052 | 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: 
22093diff
changeset | 1053 | SOME n => n | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1054 | | NONE => minsize | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1055 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1056 | 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: 
22093diff
changeset | 1057 | end; | 
| 14807 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1058 | |
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1059 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1060 | (* 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: 
14604diff
changeset | 1061 | (* 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: 
14604diff
changeset | 1062 | (* '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: 
14604diff
changeset | 1063 | (* type may have a fixed size given in 'sizes' *) | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 1064 | (* ------------------------------------------------------------------------- *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 1065 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1066 | (* (Term.typ * int) list -> (string * int) list -> int -> int -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1067 | (Term.typ * int) list option *) | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 1068 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1069 | fun next_universe xs sizes minsize maxsize = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1070 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1071 | (* 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: 
22093diff
changeset | 1072 | (* 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: 
22093diff
changeset | 1073 | (* int -> int -> int -> int list option *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1074 | fun make_first _ 0 sum = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1075 | if sum=0 then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1076 | SOME [] | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1077 | else | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1078 | NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1079 | | make_first max len sum = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1080 | if sum<=max orelse max<0 then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1081 | Option.map (fn xs' => sum :: xs') (make_first max (len-1) 0) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1082 | else | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1083 | Option.map (fn xs' => max :: xs') (make_first max (len-1) (sum-max)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1084 | (* 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: 
22093diff
changeset | 1085 | (* all list elements x (unless 'max'<0) *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1086 | (* int -> int -> int -> int list -> int list option *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1087 | fun next max len sum [] = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1088 | NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1089 | | next max len sum [x] = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1090 | (* we've reached the last list element, so there's no shift possible *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1091 | make_first max (len+1) (sum+x+1) (* increment 'sum' by 1 *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1092 | | next max len sum (x1::x2::xs) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1093 | if x1>0 andalso (x2<max orelse max<0) then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1094 | (* we can shift *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1095 | SOME (valOf (make_first max (len+1) (sum+x1-1)) @ (x2+1) :: xs) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1096 | else | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1097 | (* continue search *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1098 | next max (len+1) (sum+x1) (x2::xs) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1099 | (* only consider those types for which the size is not fixed *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1100 | val mutables = List.filter | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1101 | (not o (AList.defined (op =) sizes) o string_of_typ o fst) xs | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1102 | (* 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: 
22093diff
changeset | 1103 | 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: 
22093diff
changeset | 1104 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1105 | 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: 
22093diff
changeset | 1106 | SOME diffs' => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1107 | (* merge with those types for which the size is fixed *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1108 | SOME (snd (foldl_map (fn (ds, (T, _)) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1109 | 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: 
22093diff
changeset | 1110 | (* return the fixed size *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1111 | SOME n => (ds, (T, n)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1112 | (* consume the head of 'ds', add 'minsize' *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1113 | | NONE => (tl ds, (T, minsize + hd ds))) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1114 | (diffs', xs))) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1115 | | NONE => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1116 | NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1117 | 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: 
14604diff
changeset | 1118 | |
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1119 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1120 | (* 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: 
14604diff
changeset | 1121 | (* 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: 
14604diff
changeset | 1122 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1123 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1124 | (* interpretation -> prop_formula *) | 
| 14807 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1125 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1126 | fun toTrue (Leaf [fm, _]) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1127 | fm | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1128 | | toTrue _ = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1129 |     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: 
14604diff
changeset | 1130 | |
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1131 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1132 | (* 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: 
14604diff
changeset | 1133 | (* 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: 
14604diff
changeset | 1134 | (* 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: 
14604diff
changeset | 1135 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1136 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1137 | (* interpretation -> prop_formula *) | 
| 14807 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1138 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1139 | fun toFalse (Leaf [_, fm]) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1140 | fm | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1141 | | toFalse _ = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1142 |     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: 
14604diff
changeset | 1143 | |
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1144 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1145 | (* 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: 
14604diff
changeset | 1146 | (* 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: 
14604diff
changeset | 1147 | (* 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: 
14604diff
changeset | 1148 | (* thy : the current theory *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1149 | (* {...}     : parameters that control the translation/model generation      *)
 | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1150 | (* 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: 
14604diff
changeset | 1151 | (* 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: 
14604diff
changeset | 1152 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1153 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1154 | (* theory -> params -> Term.term -> bool -> unit *) | 
| 14807 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1155 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1156 |   fun find_model thy {sizes, minsize, maxsize, maxvars, maxtime, satsolver} t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1157 | negate = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1158 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1159 | (* unit -> unit *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1160 | fun wrapper () = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1161 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1162 | val u = unfold_defs thy t | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1163 |       val _      = writeln ("Unfolded term: " ^ Sign.string_of_term thy u)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1164 | val axioms = collect_axioms thy u | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1165 | (* Term.typ list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1166 | val types = Library.foldl (fn (acc, t') => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1167 | acc union (ground_types thy t')) ([], u :: axioms) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1168 |       val _     = writeln ("Ground types: "
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1169 | ^ (if null types then "none." | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1170 | else commas (map (Sign.string_of_typ thy) types))) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1171 | (* we can only consider fragments of recursive IDTs, so we issue a *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1172 | (* warning if the formula contains a recursive IDT *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1173 | (* TODO: no warning needed for /positive/ occurrences of IDTs *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1174 | val _ = if Library.exists (fn | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1175 | Type (s, _) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1176 | (case DatatypePackage.get_datatype thy s of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1177 | SOME info => (* inductive datatype *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1178 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1179 | val index = #index info | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1180 | val descr = #descr info | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1181 | val (_, _, constrs) = lookup descr index | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1182 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1183 | (* recursive datatype? *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1184 | Library.exists (fn (_, ds) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1185 | Library.exists DatatypeAux.is_rec_type ds) constrs | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1186 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1187 | | NONE => false) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1188 | | _ => false) types then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1189 |           warning ("Term contains a recursive datatype; "
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1190 | ^ "countermodel(s) may be spurious!") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1191 | else | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1192 | () | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1193 | (* (Term.typ * int) list -> unit *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1194 | fun find_model_loop universe = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1195 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1196 | val init_model = (universe, []) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1197 |         val init_args  = {maxvars = maxvars, def_eq = false, next_idx = 1,
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1198 | bounds = [], wellformed = True} | 
| 22580 | 1199 |         val _          = Output.immediate_output ("Translating term (sizes: "
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1200 | ^ commas (map (fn (_, n) => string_of_int n) universe) ^ ") ...") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1201 | (* translate 'u' and all axioms *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1202 | val ((model, args), intrs) = foldl_map (fn ((m, a), t') => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1203 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1204 | val (i, m', a') = interpret thy m a t' | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1205 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1206 | (* set 'def_eq' to 'true' *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1207 |             ((m', {maxvars = #maxvars a', def_eq = true,
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1208 | next_idx = #next_idx a', bounds = #bounds a', | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1209 | wellformed = #wellformed a'}), i) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1210 | end) ((init_model, init_args), u :: axioms) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1211 | (* make 'u' either true or false, and make all axioms true, and *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1212 | (* add the well-formedness side condition *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1213 | val fm_u = (if negate then toFalse else toTrue) (hd intrs) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1214 | val fm_ax = PropLogic.all (map toTrue (tl intrs)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1215 | val fm = PropLogic.all [#wellformed args, fm_ax, fm_u] | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1216 | in | 
| 22580 | 1217 | Output.immediate_output " invoking SAT solver..."; | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1218 | (case SatSolver.invoke_solver satsolver fm of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1219 | SatSolver.SATISFIABLE assignment => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1220 | (writeln " model found!"; | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1221 |           writeln ("*** Model found: ***\n" ^ print_model thy model
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1222 | (fn i => case assignment i of SOME b => b | NONE => true))) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1223 | | SatSolver.UNSATISFIABLE _ => | 
| 22580 | 1224 | (Output.immediate_output " no model exists.\n"; | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1225 | case next_universe universe sizes minsize maxsize of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1226 | SOME universe' => find_model_loop universe' | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1227 | | NONE => writeln | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1228 | "Search terminated, no larger universe within the given limits.") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1229 | | SatSolver.UNKNOWN => | 
| 22580 | 1230 | (Output.immediate_output " no model found.\n"; | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1231 | case next_universe universe sizes minsize maxsize of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1232 | SOME universe' => find_model_loop universe' | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1233 | | NONE => writeln | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1234 | "Search terminated, no larger universe within the given limits.") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1235 | ) handle SatSolver.NOT_CONFIGURED => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1236 |           error ("SAT solver " ^ quote satsolver ^ " is not configured.")
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1237 | end handle MAXVARS_EXCEEDED => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1238 |         writeln ("\nSearch terminated, number of Boolean variables ("
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1239 | ^ string_of_int maxvars ^ " allowed) exceeded.") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1240 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1241 | find_model_loop (first_universe types sizes minsize) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1242 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1243 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1244 | (* some parameter sanity checks *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1245 | minsize>=1 orelse | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1246 |         error ("\"minsize\" is " ^ string_of_int minsize ^ ", must be at least 1");
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1247 | maxsize>=1 orelse | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1248 |         error ("\"maxsize\" is " ^ string_of_int maxsize ^ ", must be at least 1");
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1249 | maxsize>=minsize orelse | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1250 |         error ("\"maxsize\" (=" ^ string_of_int maxsize ^
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1251 | ") is less than \"minsize\" (=" ^ string_of_int minsize ^ ")."); | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1252 | maxvars>=0 orelse | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1253 |         error ("\"maxvars\" is " ^ string_of_int maxvars ^ ", must be at least 0");
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1254 | maxtime>=0 orelse | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1255 |         error ("\"maxtime\" is " ^ string_of_int maxtime ^ ", must be at least 0");
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1256 | (* enter loop with or without time limit *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1257 |       writeln ("Trying to find a model that "
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1258 | ^ (if negate then "refutes" else "satisfies") ^ ": " | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1259 | ^ Sign.string_of_term thy t); | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1260 | if maxtime>0 then ( | 
| 24688 
a5754ca5c510
replaced interrupt_timeout by TimeLimit.timeLimit (available on SML/NJ and Poly/ML 5.1);
 wenzelm parents: 
23881diff
changeset | 1261 | TimeLimit.timeLimit (Time.fromSeconds maxtime) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1262 | wrapper () | 
| 24688 
a5754ca5c510
replaced interrupt_timeout by TimeLimit.timeLimit (available on SML/NJ and Poly/ML 5.1);
 wenzelm parents: 
23881diff
changeset | 1263 | handle TimeLimit.TimeOut => | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1264 |           writeln ("\nSearch terminated, time limit (" ^ string_of_int maxtime
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1265 | ^ (if maxtime=1 then " second" else " seconds") ^ ") exceeded.") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1266 | ) else | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1267 | wrapper () | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1268 | end; | 
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 1269 | |
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 1270 | |
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 1271 | (* ------------------------------------------------------------------------- *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 1272 | (* INTERFACE, PART 2: FINDING A MODEL *) | 
| 14350 | 1273 | (* ------------------------------------------------------------------------- *) | 
| 1274 | ||
| 1275 | (* ------------------------------------------------------------------------- *) | |
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 1276 | (* 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: 
14351diff
changeset | 1277 | (* params : list of '(name, value)' pairs used to override default *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 1278 | (* parameters *) | 
| 14350 | 1279 | (* ------------------------------------------------------------------------- *) | 
| 1280 | ||
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1281 | (* theory -> (string * string) list -> Term.term -> unit *) | 
| 14350 | 1282 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1283 | fun satisfy_term thy params t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1284 | find_model thy (actual_params thy params) t false; | 
| 14350 | 1285 | |
| 1286 | (* ------------------------------------------------------------------------- *) | |
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 1287 | (* 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: 
14351diff
changeset | 1288 | (* params : list of '(name, value)' pairs used to override default *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 1289 | (* parameters *) | 
| 14350 | 1290 | (* ------------------------------------------------------------------------- *) | 
| 1291 | ||
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1292 | (* theory -> (string * string) list -> Term.term -> unit *) | 
| 14350 | 1293 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1294 | fun refute_term thy params t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1295 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1296 | (* 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: 
22093diff
changeset | 1297 | (* 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: 
22093diff
changeset | 1298 | (* 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: 
22093diff
changeset | 1299 | (* for ALL types, not ...) *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1300 | val _ = null (term_tvars t) orelse | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1301 | error "Term to be refuted contains schematic type variables" | 
| 21556 | 1302 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1303 | (* existential closure over schematic variables *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1304 | (* (Term.indexname * Term.typ) list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1305 | val vars = sort_wrt (fst o fst) (map dest_Var (term_vars t)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1306 | (* Term.term *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1307 | val ex_closure = Library.foldl (fn (t', ((x, i), T)) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1308 | (HOLogic.exists_const T) $ | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1309 | Abs (x, T, abstract_over (Var ((x, i), T), t'))) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1310 | (t, vars) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1311 | (* 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: 
22093diff
changeset | 1312 | (* '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: 
22093diff
changeset | 1313 | (* 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: 
22093diff
changeset | 1314 | (* resulting term correctly, without checking its type. *) | 
| 21556 | 1315 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1316 | (* 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: 
22093diff
changeset | 1317 | (* 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: 
22093diff
changeset | 1318 | (* 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: 
22093diff
changeset | 1319 | (* 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: 
22093diff
changeset | 1320 | (* 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: 
22093diff
changeset | 1321 | (* interpretation which includes values for the (formerly) *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1322 | (* quantified variables. *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1323 | (* maps !!x1...xn. !xk...xm. t to t *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1324 |     fun strip_all_body (Const ("all", _) $ Abs (_, _, t)) = strip_all_body t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1325 |       | strip_all_body (Const ("Trueprop", _) $ t)        = strip_all_body t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1326 |       | strip_all_body (Const ("All", _) $ Abs (_, _, t)) = strip_all_body t
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1327 | | strip_all_body t = t | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1328 | (* maps !!x1...xn. !xk...xm. t to [x1, ..., xn, xk, ..., xm] *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1329 |     fun strip_all_vars (Const ("all", _) $ Abs (a, T, t)) =
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1330 | (a, T) :: strip_all_vars t | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1331 |       | strip_all_vars (Const ("Trueprop", _) $ t)        =
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1332 | strip_all_vars t | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1333 |       | strip_all_vars (Const ("All", _) $ Abs (a, T, t)) =
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1334 | (a, T) :: strip_all_vars t | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1335 | | strip_all_vars t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1336 | [] : (string * typ) list | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1337 | val strip_t = strip_all_body ex_closure | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1338 | val frees = Term.rename_wrt_term strip_t (strip_all_vars ex_closure) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1339 | 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: 
22093diff
changeset | 1340 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1341 | find_model thy (actual_params thy params) subst_t true | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1342 | end; | 
| 14350 | 1343 | |
| 1344 | (* ------------------------------------------------------------------------- *) | |
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 1345 | (* refute_subgoal: calls 'refute_term' on a specific subgoal *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 1346 | (* params : list of '(name, value)' pairs used to override default *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 1347 | (* parameters *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 1348 | (* subgoal : 0-based index specifying the subgoal number *) | 
| 14350 | 1349 | (* ------------------------------------------------------------------------- *) | 
| 1350 | ||
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1351 | (* theory -> (string * string) list -> Thm.thm -> int -> unit *) | 
| 14350 | 1352 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1353 | fun refute_subgoal thy params thm subgoal = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1354 | refute_term thy params (List.nth (Thm.prems_of thm, subgoal)); | 
| 14350 | 1355 | |
| 1356 | ||
| 1357 | (* ------------------------------------------------------------------------- *) | |
| 15292 | 1358 | (* INTERPRETERS: Auxiliary Functions *) | 
| 14350 | 1359 | (* ------------------------------------------------------------------------- *) | 
| 1360 | ||
| 1361 | (* ------------------------------------------------------------------------- *) | |
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1362 | (* make_constants: returns all interpretations for type 'T' that consist of *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1363 | (* unit vectors with 'True'/'False' only (no Boolean *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1364 | (* variables) *) | 
| 14350 | 1365 | (* ------------------------------------------------------------------------- *) | 
| 1366 | ||
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1367 | (* theory -> model -> Term.typ -> interpretation list *) | 
| 14350 | 1368 | |
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1369 | fun make_constants thy model T = | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1370 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1371 | (* 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: 
22093diff
changeset | 1372 | (* int -> interpretation list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1373 | fun unit_vectors n = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1374 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1375 | (* returns the k-th unit vector of length n *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1376 | (* int * int -> interpretation *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1377 | fun unit_vector (k, n) = | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1378 | Leaf ((replicate (k-1) False) @ (True :: (replicate (n-k) False))) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1379 | (* int -> interpretation list *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1380 | fun unit_vectors_loop k = | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1381 | if k>n then [] else unit_vector (k,n) :: unit_vectors_loop (k+1) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1382 | in | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1383 | unit_vectors_loop 1 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1384 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1385 | (* 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: 
22093diff
changeset | 1386 | (* identical) elements from 'xs' *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1387 | (* int -> 'a list -> 'a list list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1388 | fun pick_all 1 xs = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1389 | map single xs | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1390 | | pick_all n xs = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1391 | let val rec_pick = pick_all (n-1) xs in | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1392 | List.concat (map (fn x => map (cons x) rec_pick) xs) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1393 | end | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1394 | (* returns all constant interpretations that have the same tree *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1395 | (* structure as the interpretation argument *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1396 | (* interpretation -> interpretation list *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1397 | fun make_constants_intr (Leaf xs) = unit_vectors (length xs) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1398 | | make_constants_intr (Node xs) = map Node (pick_all (length xs) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1399 | (make_constants_intr (hd xs))) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1400 | (* obtain the interpretation for a variable of type 'T' *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1401 |     val (i, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1,
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1402 |       bounds=[], wellformed=True} (Free ("dummy", T))
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1403 | in | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1404 | make_constants_intr i | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1405 | 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: 
14604diff
changeset | 1406 | |
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1407 | (* ------------------------------------------------------------------------- *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1408 | (* power: 'power (a, b)' computes a^b, for a>=0, b>=0 *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1409 | (* ------------------------------------------------------------------------- *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1410 | |
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1411 | (* int * int -> int *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1412 | |
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1413 | fun power (a, 0) = 1 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1414 | | power (a, 1) = a | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1415 | | power (a, b) = let val ab = power(a, b div 2) in | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1416 | ab * ab * power(a, b mod 2) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1417 | end; | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1418 | |
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1419 | (* ------------------------------------------------------------------------- *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1420 | (* 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: 
24928diff
changeset | 1421 | (* (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: 
14604diff
changeset | 1422 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1423 | |
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1424 | (* theory -> model -> Term.typ -> int *) | 
| 14807 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1425 | |
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1426 | (* returns 0 for an empty ground type or a function type with empty *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1427 | (* codomain, but fails for a function type with empty domain -- *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1428 | (* admissibility of datatype constructor argument types (see "Inductive *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1429 | (* datatypes in HOL - lessons learned ...", S. Berghofer, M. Wenzel, *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1430 | (* TPHOLs 99) ensures that recursive, possibly empty, datatype fragments *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1431 | (* never occur as the domain of a function type that is the type of a *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1432 | (* constructor argument *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1433 | |
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1434 | fun size_of_type thy model T = | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1435 | let | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1436 | (* returns the number of elements that have the same tree structure as a *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1437 | (* given interpretation *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1438 | fun size_of_intr (Leaf xs) = length xs | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1439 | | size_of_intr (Node xs) = power (size_of_intr (hd xs), length xs) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1440 | (* obtain the interpretation for a variable of type 'T' *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1441 |     val (i, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1,
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1442 |       bounds=[], wellformed=True} (Free ("dummy", T))
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1443 | in | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1444 | size_of_intr i | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1445 | 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: 
14604diff
changeset | 1446 | |
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1447 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1448 | (* 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: 
14604diff
changeset | 1449 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1450 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1451 | (* 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: 
14604diff
changeset | 1452 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1453 | 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: 
14604diff
changeset | 1454 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1455 | 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: 
14604diff
changeset | 1456 | |
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1457 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1458 | (* 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: 
14604diff
changeset | 1459 | (* 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: 
15531diff
changeset | 1460 | (* - 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: 
15531diff
changeset | 1461 | (* 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: 
15531diff
changeset | 1462 | (* - 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: 
15531diff
changeset | 1463 | (* 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: 
15531diff
changeset | 1464 | (* - 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: 
15531diff
changeset | 1465 | (* '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: 
14604diff
changeset | 1466 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1467 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1468 | (* We could in principle represent '=' on a type T by a particular *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1469 | (* interpretation. However, the size of that interpretation is quadratic *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1470 | (* in the size of T. Therefore comparing the interpretations 'i1' and *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1471 | (* 'i2' directly is more efficient than constructing the interpretation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1472 | (* for equality on T first, and "applying" this interpretation to 'i1' *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1473 | (* 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: 
14604diff
changeset | 1474 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1475 | (* interpretation * interpretation -> interpretation *) | 
| 14807 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1476 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1477 | fun make_equality (i1, i2) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1478 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1479 | (* interpretation * interpretation -> prop_formula *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1480 | fun equal (i1, i2) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1481 | (case i1 of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1482 | Leaf xs => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1483 | (case i2 of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1484 | Leaf ys => PropLogic.dot_product (xs, ys) (* defined and equal *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1485 |         | Node _  => raise REFUTE ("make_equality",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1486 | "second interpretation is higher")) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1487 | | Node xs => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1488 | (case i2 of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1489 |           Leaf _  => raise REFUTE ("make_equality",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1490 | "first interpretation is higher") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1491 | | Node ys => PropLogic.all (map equal (xs ~~ ys)))) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1492 | (* interpretation * interpretation -> prop_formula *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1493 | fun not_equal (i1, i2) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1494 | (case i1 of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1495 | Leaf xs => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1496 | (case i2 of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1497 | (* defined and not equal *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1498 | Leaf ys => PropLogic.all ((PropLogic.exists xs) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1499 | :: (PropLogic.exists ys) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1500 | :: (map (fn (x,y) => SOr (SNot x, SNot y)) (xs ~~ ys))) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1501 |         | Node _  => raise REFUTE ("make_equality",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1502 | "second interpretation is higher")) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1503 | | Node xs => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1504 | (case i2 of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1505 |           Leaf _  => raise REFUTE ("make_equality",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1506 | "first interpretation is higher") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1507 | | Node ys => PropLogic.exists (map not_equal (xs ~~ ys)))) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1508 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1509 | (* 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: 
22093diff
changeset | 1510 | (* negation of 'equal' *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1511 | 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: 
22093diff
changeset | 1512 | 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: 
14604diff
changeset | 1513 | |
| 15292 | 1514 | (* ------------------------------------------------------------------------- *) | 
| 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: 
15531diff
changeset | 1515 | (* 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: 
15531diff
changeset | 1516 | (* 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: 
15531diff
changeset | 1517 | (* 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: 
15531diff
changeset | 1518 | (* 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: 
15531diff
changeset | 1519 | (* 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: 
15531diff
changeset | 1520 | (* 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: 
15531diff
changeset | 1521 | (* ------------------------------------------------------------------------- *) | 
| 
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: 
15531diff
changeset | 1522 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1523 | (* interpretation * interpretation -> interpretation *) | 
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15531diff
changeset | 1524 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1525 | fun make_def_equality (i1, i2) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1526 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1527 | (* interpretation * interpretation -> prop_formula *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1528 | fun equal (i1, i2) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1529 | (case i1 of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1530 | Leaf xs => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1531 | (case i2 of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1532 | (* defined and equal, or both undefined *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1533 | Leaf ys => SOr (PropLogic.dot_product (xs, ys), | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1534 | SAnd (PropLogic.all (map SNot xs), PropLogic.all (map SNot ys))) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1535 |         | Node _  => raise REFUTE ("make_def_equality",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1536 | "second interpretation is higher")) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1537 | | Node xs => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1538 | (case i2 of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1539 |           Leaf _  => raise REFUTE ("make_def_equality",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1540 | "first interpretation is higher") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1541 | | Node ys => PropLogic.all (map equal (xs ~~ ys)))) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1542 | (* interpretation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1543 | val eq = equal (i1, i2) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1544 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1545 | Leaf [eq, SNot eq] | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1546 | 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: 
15531diff
changeset | 1547 | |
| 
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: 
15531diff
changeset | 1548 | (* ------------------------------------------------------------------------- *) | 
| 
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: 
15531diff
changeset | 1549 | (* interpretation_apply: returns an interpretation that denotes the result *) | 
| 22092 | 1550 | (* 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: 
15531diff
changeset | 1551 | (* 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: 
15531diff
changeset | 1552 | (* ------------------------------------------------------------------------- *) | 
| 
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: 
15531diff
changeset | 1553 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1554 | (* interpretation * interpretation -> interpretation *) | 
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15531diff
changeset | 1555 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1556 | fun interpretation_apply (i1, i2) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1557 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1558 | (* interpretation * interpretation -> interpretation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1559 | fun interpretation_disjunction (tr1,tr2) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1560 | 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: 
22093diff
changeset | 1561 | (tree_pair (tr1,tr2)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1562 | (* prop_formula * interpretation -> interpretation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1563 | fun prop_formula_times_interpretation (fm,tr) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1564 | 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: 
22093diff
changeset | 1565 | (* prop_formula list * interpretation list -> interpretation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1566 | fun prop_formula_list_dot_product_interpretation_list ([fm],[tr]) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1567 | prop_formula_times_interpretation (fm,tr) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1568 | | prop_formula_list_dot_product_interpretation_list (fm::fms,tr::trees) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1569 | interpretation_disjunction (prop_formula_times_interpretation (fm,tr), | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1570 | prop_formula_list_dot_product_interpretation_list (fms,trees)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1571 | | prop_formula_list_dot_product_interpretation_list (_,_) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1572 |       raise REFUTE ("interpretation_apply", "empty list (in dot product)")
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1573 | (* concatenates 'x' with every list in 'xss', returning a new list of *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1574 | (* lists *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1575 | (* 'a -> 'a list list -> 'a list list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1576 | fun cons_list x xss = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1577 | map (cons x) xss | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1578 | (* 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: 
22093diff
changeset | 1579 | (* element of 'xss' *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1580 | (* 'a list list -> 'a list list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1581 | fun pick_all [xs] = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1582 | map single xs | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1583 | | pick_all (xs::xss) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1584 | let val rec_pick = pick_all xss in | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1585 | List.concat (map (fn x => map (cons x) rec_pick) xs) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1586 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1587 | | pick_all _ = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1588 |       raise REFUTE ("interpretation_apply", "empty list (in pick_all)")
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1589 | (* interpretation -> prop_formula list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1590 | fun interpretation_to_prop_formula_list (Leaf xs) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1591 | xs | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1592 | | interpretation_to_prop_formula_list (Node trees) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1593 | map PropLogic.all (pick_all | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1594 | (map interpretation_to_prop_formula_list trees)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1595 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1596 | case i1 of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1597 | Leaf _ => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1598 |       raise REFUTE ("interpretation_apply", "first interpretation is a leaf")
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1599 | | Node xs => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1600 | prop_formula_list_dot_product_interpretation_list | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1601 | (interpretation_to_prop_formula_list i2, xs) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1602 | 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: 
15531diff
changeset | 1603 | |
| 
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: 
15531diff
changeset | 1604 | (* ------------------------------------------------------------------------- *) | 
| 15292 | 1605 | (* eta_expand: eta-expands a term 't' by adding 'i' lambda abstractions *) | 
| 1606 | (* ------------------------------------------------------------------------- *) | |
| 1607 | ||
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1608 | (* Term.term -> int -> Term.term *) | 
| 15292 | 1609 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1610 | fun eta_expand t i = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1611 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1612 | 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: 
22093diff
changeset | 1613 | val t' = Term.incr_boundvars i t | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1614 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1615 |     foldr (fn (T, term) => Abs ("<eta_expand>", T, term))
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1616 | (Term.list_comb (t', map Bound (i-1 downto 0))) (List.take (Ts, i)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1617 | end; | 
| 15292 | 1618 | |
| 15335 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 1619 | (* ------------------------------------------------------------------------- *) | 
| 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 1620 | (* sum: returns the sum of a list 'xs' of integers *) | 
| 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 1621 | (* ------------------------------------------------------------------------- *) | 
| 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 1622 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1623 | (* int list -> int *) | 
| 15335 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 1624 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1625 | fun sum xs = foldl op+ 0 xs; | 
| 15335 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 1626 | |
| 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 1627 | (* ------------------------------------------------------------------------- *) | 
| 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 1628 | (* product: returns the product of a list 'xs' of integers *) | 
| 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 1629 | (* ------------------------------------------------------------------------- *) | 
| 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 1630 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1631 | (* int list -> int *) | 
| 15335 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 1632 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1633 | fun product xs = foldl op* 1 xs; | 
| 15335 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 1634 | |
| 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 1635 | (* ------------------------------------------------------------------------- *) | 
| 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: 
15531diff
changeset | 1636 | (* 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: 
15531diff
changeset | 1637 | (* 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: 
15531diff
changeset | 1638 | (* their arguments) of the size of the argument types *) | 
| 15335 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 1639 | (* ------------------------------------------------------------------------- *) | 
| 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 1640 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1641 | (* theory -> (Term.typ * int) list -> DatatypeAux.descr -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1642 | (DatatypeAux.dtyp * Term.typ) list -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1643 | (string * DatatypeAux.dtyp list) list -> int *) | 
| 15335 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 1644 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1645 | fun size_of_dtyp thy typ_sizes descr typ_assoc constructors = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1646 | sum (map (fn (_, dtyps) => | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1647 | product (map (size_of_type thy (typ_sizes, []) o | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1648 | (typ_of_dtyp descr typ_assoc)) dtyps)) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1649 | constructors); | 
| 15335 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 webertj parents: 
15334diff
changeset | 1650 | |
| 15292 | 1651 | |
| 1652 | (* ------------------------------------------------------------------------- *) | |
| 1653 | (* INTERPRETERS: Actual Interpreters *) | |
| 1654 | (* ------------------------------------------------------------------------- *) | |
| 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: 
14604diff
changeset | 1655 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1656 | (* theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1657 | (interpretation * model * arguments) option *) | 
| 14807 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1658 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1659 | (* simply typed lambda calculus: Isabelle's basic term syntax, with type *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1660 | (* 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: 
14604diff
changeset | 1661 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1662 | fun stlc_interpreter thy model args t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1663 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1664 | val (typs, terms) = model | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1665 |     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: 
22093diff
changeset | 1666 | (* Term.typ -> (interpretation * model * arguments) option *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1667 | fun interpret_groundterm T = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1668 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1669 | (* unit -> (interpretation * model * arguments) option *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1670 | fun interpret_groundtype () = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1671 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1672 | (* the model must specify a size for ground types *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1673 | val size = (if T = Term.propT then 2 else lookup typs T) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1674 | val next = next_idx+size | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1675 | (* check if 'maxvars' is large enough *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1676 | val _ = (if next-1>maxvars andalso maxvars>0 then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1677 | raise MAXVARS_EXCEEDED else ()) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1678 | (* prop_formula list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1679 | val fms = map BoolVar (next_idx upto (next_idx+size-1)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1680 | (* interpretation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1681 | val intr = Leaf fms | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1682 | (* prop_formula list -> prop_formula *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1683 | fun one_of_two_false [] = True | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1684 | | one_of_two_false (x::xs) = SAnd (PropLogic.all (map (fn x' => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1685 | SOr (SNot x, SNot x')) xs), one_of_two_false xs) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1686 | (* prop_formula *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1687 | val wf = one_of_two_false fms | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1688 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1689 | (* extend the model, increase 'next_idx', add well-formedness *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1690 | (* condition *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1691 |         SOME (intr, (typs, (t, intr)::terms), {maxvars = maxvars,
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1692 | def_eq = def_eq, next_idx = next, bounds = bounds, | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1693 | wellformed = SAnd (wellformed, wf)}) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1694 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1695 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1696 | case T of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1697 |         Type ("fun", [T1, T2]) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1698 | let | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1699 | (* we create 'size_of_type ... T1' different copies of the *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1700 | (* interpretation for 'T2', which are then combined into a single *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1701 | (* new interpretation *) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1702 | (* make fresh copies, with different variable indices *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1703 | (* 'idx': next variable index *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1704 | (* 'n' : number of copies *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1705 | (* int -> int -> (int * interpretation list * prop_formula *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1706 | fun make_copies idx 0 = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1707 | (idx, [], True) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1708 | | make_copies idx n = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1709 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1710 | val (copy, _, new_args) = interpret thy (typs, []) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1711 |                 {maxvars = maxvars, def_eq = false, next_idx = idx,
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1712 |                 bounds = [], wellformed = True} (Free ("dummy", T2))
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1713 | val (idx', copies, wf') = make_copies (#next_idx new_args) (n-1) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1714 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1715 | (idx', copy :: copies, SAnd (#wellformed new_args, wf')) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1716 | end | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1717 | val (next, copies, wf) = make_copies next_idx | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1718 | (size_of_type thy model T1) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1719 | (* combine copies into a single interpretation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1720 | val intr = Node copies | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1721 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1722 | (* extend the model, increase 'next_idx', add well-formedness *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1723 | (* condition *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1724 |           SOME (intr, (typs, (t, intr)::terms), {maxvars = maxvars,
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1725 | def_eq = def_eq, next_idx = next, bounds = bounds, | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1726 | wellformed = SAnd (wellformed, wf)}) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1727 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1728 | | Type _ => interpret_groundtype () | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1729 | | TFree _ => interpret_groundtype () | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1730 | | TVar _ => interpret_groundtype () | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1731 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1732 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1733 | case AList.lookup (op =) terms t of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1734 | SOME intr => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1735 | (* return an existing interpretation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1736 | SOME (intr, model, args) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1737 | | NONE => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1738 | (case t of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1739 | Const (_, T) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1740 | interpret_groundterm T | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1741 | | Free (_, T) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1742 | interpret_groundterm T | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1743 | | Var (_, T) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1744 | interpret_groundterm T | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1745 | | Bound i => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1746 | SOME (List.nth (#bounds args, i), model, args) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1747 | | Abs (x, T, body) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1748 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1749 | (* create all constants of type 'T' *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 1750 | val constants = make_constants thy model T | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1751 | (* interpret the 'body' separately for each constant *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1752 | val ((model', args'), bodies) = foldl_map | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1753 | (fn ((m, a), c) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1754 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1755 | (* add 'c' to 'bounds' *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1756 |                 val (i', m', a') = interpret thy m {maxvars = #maxvars a,
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1757 | def_eq = #def_eq a, next_idx = #next_idx a, | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1758 | bounds = (c :: #bounds a), wellformed = #wellformed a} body | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1759 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1760 | (* keep the new model m' and 'next_idx' and 'wellformed', *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1761 | (* but use old 'bounds' *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1762 |                 ((m', {maxvars = maxvars, def_eq = def_eq,
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1763 | next_idx = #next_idx a', bounds = bounds, | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1764 | wellformed = #wellformed a'}), i') | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1765 | end) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1766 | ((model, args), constants) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1767 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1768 | SOME (Node bodies, model', args') | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1769 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1770 | | t1 $ t2 => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1771 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1772 | (* interpret 't1' and 't2' separately *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1773 | val (intr1, model1, args1) = interpret thy model args t1 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1774 | val (intr2, model2, args2) = interpret thy model1 args1 t2 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1775 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1776 | SOME (interpretation_apply (intr1, intr2), model2, args2) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1777 | end) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1778 | 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: 
14604diff
changeset | 1779 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1780 | (* theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1781 | (interpretation * model * arguments) option *) | 
| 14807 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1782 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1783 | fun Pure_interpreter thy model args t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1784 | case t of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1785 |       Const ("all", _) $ t1 =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1786 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1787 | val (i, m, a) = interpret thy model args t1 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1788 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1789 | case i of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1790 | Node xs => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1791 | (* 3-valued logic *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1792 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1793 | val fmTrue = PropLogic.all (map toTrue xs) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1794 | val fmFalse = PropLogic.exists (map toFalse xs) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1795 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1796 | SOME (Leaf [fmTrue, fmFalse], m, a) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1797 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1798 | | _ => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1799 |           raise REFUTE ("Pure_interpreter",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1800 | "\"all\" is followed by a non-function") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1801 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1802 |     | Const ("all", _) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1803 | SOME (interpret thy model args (eta_expand t 1)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1804 |     | Const ("==", _) $ t1 $ t2 =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1805 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1806 | val (i1, m1, a1) = interpret thy model args t1 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1807 | val (i2, m2, a2) = interpret thy m1 a1 t2 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1808 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1809 | (* 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: 
22093diff
changeset | 1810 | 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: 
22093diff
changeset | 1811 | (i1, i2), m2, a2) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1812 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1813 |     | Const ("==", _) $ t1 =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1814 | SOME (interpret thy model args (eta_expand t 1)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1815 |     | Const ("==", _) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1816 | SOME (interpret thy model args (eta_expand t 2)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1817 |     | Const ("==>", _) $ t1 $ t2 =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1818 | (* 3-valued logic *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1819 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1820 | val (i1, m1, a1) = interpret thy model args t1 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1821 | val (i2, m2, a2) = interpret thy m1 a1 t2 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1822 | val fmTrue = PropLogic.SOr (toFalse i1, toTrue i2) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1823 | val fmFalse = PropLogic.SAnd (toTrue i1, toFalse i2) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1824 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1825 | SOME (Leaf [fmTrue, fmFalse], m2, a2) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1826 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1827 |     | Const ("==>", _) $ t1 =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1828 | SOME (interpret thy model args (eta_expand t 1)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1829 |     | Const ("==>", _) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1830 | SOME (interpret thy model args (eta_expand t 2)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1831 | | _ => 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: 
14604diff
changeset | 1832 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1833 | (* theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1834 | (interpretation * model * arguments) option *) | 
| 14807 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1835 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1836 | fun HOLogic_interpreter thy model args t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1837 | (* Providing interpretations directly is more efficient than unfolding the *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1838 | (* logical constants. In HOL however, logical constants can themselves be *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1839 | (* arguments. They are then translated using eta-expansion. *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1840 | case t of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1841 |       Const ("Trueprop", _) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1842 | SOME (Node [TT, FF], model, args) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1843 |     | Const ("Not", _) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1844 | SOME (Node [FF, TT], model, args) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1845 | (* redundant, since 'True' is also an IDT constructor *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1846 |     | Const ("True", _) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1847 | SOME (TT, model, args) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1848 | (* redundant, since 'False' is also an IDT constructor *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1849 |     | Const ("False", _) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1850 | SOME (FF, model, args) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1851 |     | Const ("All", _) $ t1 =>  (* similar to "all" (Pure) *)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1852 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1853 | val (i, m, a) = interpret thy model args t1 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1854 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1855 | case i of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1856 | Node xs => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1857 | (* 3-valued logic *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1858 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1859 | val fmTrue = PropLogic.all (map toTrue xs) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1860 | val fmFalse = PropLogic.exists (map toFalse xs) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1861 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1862 | SOME (Leaf [fmTrue, fmFalse], m, a) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1863 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1864 | | _ => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1865 |           raise REFUTE ("HOLogic_interpreter",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1866 | "\"All\" is followed by a non-function") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1867 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1868 |     | Const ("All", _) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1869 | SOME (interpret thy model args (eta_expand t 1)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1870 |     | Const ("Ex", _) $ t1 =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1871 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1872 | val (i, m, a) = interpret thy model args t1 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1873 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1874 | case i of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1875 | Node xs => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1876 | (* 3-valued logic *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1877 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1878 | val fmTrue = PropLogic.exists (map toTrue xs) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1879 | val fmFalse = PropLogic.all (map toFalse xs) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1880 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1881 | SOME (Leaf [fmTrue, fmFalse], m, a) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1882 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1883 | | _ => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1884 |           raise REFUTE ("HOLogic_interpreter",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1885 | "\"Ex\" is followed by a non-function") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1886 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1887 |     | Const ("Ex", _) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1888 | SOME (interpret thy model args (eta_expand t 1)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1889 |     | Const ("op =", _) $ t1 $ t2 =>  (* similar to "==" (Pure) *)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1890 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1891 | val (i1, m1, a1) = interpret thy model args t1 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1892 | val (i2, m2, a2) = interpret thy m1 a1 t2 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1893 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1894 | SOME (make_equality (i1, i2), m2, a2) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1895 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1896 |     | Const ("op =", _) $ t1 =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1897 | SOME (interpret thy model args (eta_expand t 1)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1898 |     | Const ("op =", _) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1899 | SOME (interpret thy model args (eta_expand t 2)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1900 |     | Const ("op &", _) $ t1 $ t2 =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1901 | (* 3-valued logic *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1902 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1903 | val (i1, m1, a1) = interpret thy model args t1 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1904 | val (i2, m2, a2) = interpret thy m1 a1 t2 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1905 | val fmTrue = PropLogic.SAnd (toTrue i1, toTrue i2) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1906 | val fmFalse = PropLogic.SOr (toFalse i1, toFalse i2) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1907 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1908 | SOME (Leaf [fmTrue, fmFalse], m2, a2) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1909 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1910 |     | Const ("op &", _) $ t1 =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1911 | SOME (interpret thy model args (eta_expand t 1)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1912 |     | Const ("op &", _) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1913 | SOME (interpret thy model args (eta_expand t 2)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1914 | (* 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: 
22093diff
changeset | 1915 | (* "False & undef": *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1916 | (* SOME (Node [Node [TT, FF], Node [FF, FF]], model, args) *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1917 |     | Const ("op |", _) $ t1 $ t2 =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1918 | (* 3-valued logic *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1919 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1920 | val (i1, m1, a1) = interpret thy model args t1 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1921 | val (i2, m2, a2) = interpret thy m1 a1 t2 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1922 | val fmTrue = PropLogic.SOr (toTrue i1, toTrue i2) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1923 | val fmFalse = PropLogic.SAnd (toFalse i1, toFalse i2) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1924 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1925 | SOME (Leaf [fmTrue, fmFalse], m2, a2) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1926 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1927 |     | Const ("op |", _) $ t1 =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1928 | SOME (interpret thy model args (eta_expand t 1)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1929 |     | Const ("op |", _) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1930 | SOME (interpret thy model args (eta_expand t 2)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1931 | (* 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: 
22093diff
changeset | 1932 | (* "True | undef": *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1933 | (* SOME (Node [Node [TT, TT], Node [TT, FF]], model, args) *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1934 |     | Const ("op -->", _) $ t1 $ t2 =>  (* similar to "==>" (Pure) *)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1935 | (* 3-valued logic *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1936 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1937 | val (i1, m1, a1) = interpret thy model args t1 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1938 | val (i2, m2, a2) = interpret thy m1 a1 t2 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1939 | val fmTrue = PropLogic.SOr (toFalse i1, toTrue i2) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1940 | val fmFalse = PropLogic.SAnd (toTrue i1, toFalse i2) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1941 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1942 | SOME (Leaf [fmTrue, fmFalse], m2, a2) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1943 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1944 |     | Const ("op -->", _) $ t1 =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1945 | SOME (interpret thy model args (eta_expand t 1)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1946 |     | Const ("op -->", _) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1947 | SOME (interpret thy model args (eta_expand t 2)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1948 | (* 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: 
22093diff
changeset | 1949 | (* "False --> undef": *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1950 | (* SOME (Node [Node [TT, FF], Node [TT, TT]], model, args) *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1951 | | _ => 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: 
14604diff
changeset | 1952 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1953 | (* theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1954 | (interpretation * model * arguments) option *) | 
| 14807 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 1955 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1956 | fun set_interpreter thy model args t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1957 | (* "T set" is isomorphic to "T --> bool" *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1958 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1959 | val (typs, terms) = model | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1960 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1961 | case AList.lookup (op =) terms t of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1962 | SOME intr => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1963 | (* return an existing interpretation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1964 | SOME (intr, model, args) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1965 | | NONE => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1966 | (case t of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1967 |         Free (x, Type ("set", [T])) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1968 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1969 | val (intr, _, args') = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1970 | interpret thy (typs, []) args (Free (x, T --> HOLogic.boolT)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1971 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1972 | SOME (intr, (typs, (t, intr)::terms), args') | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1973 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1974 |       | Var ((x, i), Type ("set", [T])) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1975 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1976 | val (intr, _, args') = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1977 | interpret thy (typs, []) args (Var ((x,i), T --> HOLogic.boolT)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1978 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1979 | SOME (intr, (typs, (t, intr)::terms), args') | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1980 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1981 |       | Const (s, Type ("set", [T])) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1982 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1983 | val (intr, _, args') = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1984 | interpret thy (typs, []) args (Const (s, T --> HOLogic.boolT)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1985 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1986 | SOME (intr, (typs, (t, intr)::terms), args') | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1987 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1988 | (* 'Collect' == identity *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1989 |       | Const ("Collect", _) $ t1 =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1990 | SOME (interpret thy model args t1) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1991 |       | Const ("Collect", _) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1992 | SOME (interpret thy model args (eta_expand t 1)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1993 | (* 'op :' == application *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1994 |       | Const ("op :", _) $ t1 $ t2 =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1995 | SOME (interpret thy model args (t2 $ t1)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1996 |       | Const ("op :", _) $ t1 =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1997 | SOME (interpret thy model args (eta_expand t 1)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1998 |       | Const ("op :", _) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 1999 | SOME (interpret thy model args (eta_expand t 2)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2000 | | _ => NONE) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2001 | 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: 
14604diff
changeset | 2002 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2003 | (* theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2004 | (interpretation * model * arguments) option *) | 
| 14807 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 2005 | |
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2006 | (* interprets variables and constants whose type is an IDT (this is *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2007 | (* relatively easy and merely requires us to compute the size of the IDT); *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2008 | (* constructors of IDTs however are properly interpreted by *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2009 | (* '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: 
15531diff
changeset | 2010 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2011 | fun IDT_interpreter thy model args t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2012 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2013 | val (typs, terms) = model | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2014 | (* Term.typ -> (interpretation * model * arguments) option *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2015 | fun interpret_term (Type (s, Ts)) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2016 | (case DatatypePackage.get_datatype thy s of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2017 | SOME info => (* inductive datatype *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2018 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2019 | (* int option -- only recursive IDTs have an associated depth *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2020 | val depth = AList.lookup (op =) typs (Type (s, Ts)) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2021 | (* sanity check: depth must be at least 0 *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2022 | val _ = (case depth of SOME n => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2023 | if n<0 then | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2024 |               raise REFUTE ("IDT_interpreter", "negative depth")
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2025 | else () | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2026 | | _ => ()) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2027 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2028 | (* termination condition to avoid infinite recursion *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2029 | if depth = (SOME 0) then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2030 | (* return a leaf of size 0 *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2031 | SOME (Leaf [], model, args) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2032 | else | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2033 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2034 | val index = #index info | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2035 | val descr = #descr info | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2036 | val (_, dtyps, constrs) = lookup descr index | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2037 | val typ_assoc = dtyps ~~ Ts | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2038 | (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2039 | val _ = if Library.exists (fn d => | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2040 | case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2041 | then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2042 |                   raise REFUTE ("IDT_interpreter",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2043 | "datatype argument (for type " | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2044 | ^ Sign.string_of_typ thy (Type (s, Ts)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2045 | ^ ") is not a variable") | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2046 | else () | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2047 | (* if the model specifies a depth for the current type, *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2048 | (* decrement it to avoid infinite recursion *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2049 | val typs' = case depth of NONE => typs | SOME n => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2050 | AList.update (op =) (Type (s, Ts), n-1) typs | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2051 | (* recursively compute the size of the datatype *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2052 | val size = size_of_dtyp thy typs' descr typ_assoc constrs | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2053 | val next_idx = #next_idx args | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2054 | val next = next_idx+size | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2055 | (* check if 'maxvars' is large enough *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2056 | val _ = (if next-1 > #maxvars args andalso | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2057 | #maxvars args > 0 then raise MAXVARS_EXCEEDED else ()) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2058 | (* prop_formula list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2059 | val fms = map BoolVar (next_idx upto (next_idx+size-1)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2060 | (* interpretation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2061 | val intr = Leaf fms | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2062 | (* prop_formula list -> prop_formula *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2063 | fun one_of_two_false [] = True | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2064 | | one_of_two_false (x::xs) = SAnd (PropLogic.all (map (fn x' => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2065 | SOr (SNot x, SNot x')) xs), one_of_two_false xs) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2066 | (* prop_formula *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2067 | val wf = one_of_two_false fms | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2068 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2069 | (* extend the model, increase 'next_idx', add well-formedness *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2070 | (* condition *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2071 |               SOME (intr, (typs, (t, intr)::terms), {maxvars = #maxvars args,
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2072 | def_eq = #def_eq args, next_idx = next, bounds = #bounds args, | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2073 | wellformed = SAnd (#wellformed args, wf)}) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2074 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2075 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2076 | | NONE => (* not an inductive datatype *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2077 | NONE) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2078 | | interpret_term _ = (* a (free or schematic) type variable *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2079 | NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2080 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2081 | case AList.lookup (op =) terms t of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2082 | SOME intr => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2083 | (* return an existing interpretation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2084 | SOME (intr, model, args) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2085 | | NONE => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2086 | (case t of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2087 | Free (_, T) => interpret_term T | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2088 | | Var (_, T) => interpret_term T | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2089 | | Const (_, T) => interpret_term T | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2090 | | _ => NONE) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2091 | 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: 
15531diff
changeset | 2092 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2093 | (* theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2094 | (interpretation * model * arguments) option *) | 
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15531diff
changeset | 2095 | |
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2096 | (* This function imposes an order on the elements of a datatype fragment *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2097 | (* as follows: C_i x_1 ... x_n < C_j y_1 ... y_m iff i < j or *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2098 | (* (x_1, ..., x_n) < (y_1, ..., y_m). With this order, a constructor is *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2099 | (* a function C_i that maps some argument indices x_1, ..., x_n to the *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2100 | (* datatype element given by index C_i x_1 ... x_n. The idea remains the *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2101 | (* same for recursive datatypes, although the computation of indices gets *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2102 | (* a little tricky. *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2103 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2104 | fun IDT_constructor_interpreter thy model args t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2105 | let | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2106 | (* returns a list of canonical representations for terms of the type 'T' *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2107 | (* 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: 
24928diff
changeset | 2108 | (* for IDTs calls 'IDT_constructor_interpreter' again, and this could *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2109 | (* lead to infinite recursion when we have (mutually) recursive IDTs. *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2110 | (* (Term.typ * int) list -> Term.typ -> Term.term list *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2111 | fun canonical_terms typs T = | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2112 | (case T of | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2113 |         Type ("fun", [T1, T2]) =>
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2114 | (* 'T2' might contain a recursive IDT, so we cannot use 'print' (at *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2115 | (* least not for 'T2' *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2116 | let | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2117 | (* returns a list of lists, each one consisting of n (possibly *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2118 | (* identical) elements from 'xs' *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2119 | (* int -> 'a list -> 'a list list *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2120 | fun pick_all 1 xs = | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2121 | map single xs | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2122 | | pick_all n xs = | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2123 | let val rec_pick = pick_all (n-1) xs in | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2124 | List.concat (map (fn x => map (cons x) rec_pick) xs) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2125 | end | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2126 | (* ["x1", ..., "xn"] *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2127 | val terms1 = canonical_terms typs T1 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2128 | (* ["y1", ..., "ym"] *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2129 | val terms2 = canonical_terms typs T2 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2130 |           (* [[("x1", "y1"), ..., ("xn", "y1")], ..., *)
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2131 |           (*   [("x1", "ym"), ..., ("xn", "ym")]]     *)
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2132 | val functions = map (curry (op ~~) terms1) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2133 | (pick_all (length terms1) terms2) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2134 | (* [["(x1, y1)", ..., "(xn, y1)"], ..., *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2135 | (* ["(x1, ym)", ..., "(xn, ym)"]] *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2136 | val pairss = map (map HOLogic.mk_prod) functions | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2137 | (* Term.typ *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2138 | val HOLogic_prodT = HOLogic.mk_prodT (T1, T2) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2139 | val HOLogic_setT = HOLogic.mk_setT HOLogic_prodT | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2140 | (* Term.term *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2141 |           val HOLogic_empty_set = Const ("{}", HOLogic_setT)
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2142 | val HOLogic_insert = | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2143 |             Const ("insert", HOLogic_prodT --> HOLogic_setT --> HOLogic_setT)
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2144 | in | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2145 | (* functions as graphs, i.e. as a (HOL) set of pairs "(x, y)" *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2146 | map (foldr (fn (pair, acc) => HOLogic_insert $ pair $ acc) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2147 | HOLogic_empty_set) pairss | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2148 | end | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2149 | | Type (s, Ts) => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2150 | (case DatatypePackage.get_datatype thy s of | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2151 | SOME info => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2152 | (case AList.lookup (op =) typs T of | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2153 | SOME 0 => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2154 | (* termination condition to avoid infinite recursion *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2155 | [] (* at depth 0, every IDT is empty *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2156 | | _ => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2157 | let | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2158 | val index = #index info | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2159 | val descr = #descr info | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2160 | val (_, dtyps, constrs) = lookup descr index | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2161 | val typ_assoc = dtyps ~~ Ts | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2162 | (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2163 | val _ = if Library.exists (fn d => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2164 | case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2165 | then | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2166 |                   raise REFUTE ("IDT_constructor_interpreter",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2167 | "datatype argument (for type " | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2168 | ^ Sign.string_of_typ thy T | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2169 | ^ ") is not a variable") | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2170 | else () | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2171 | (* decrement depth for the IDT 'T' *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2172 | val typs' = (case AList.lookup (op =) typs T of NONE => typs | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2173 | | SOME n => AList.update (op =) (T, n-1) typs) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2174 | (* Term.term list -> DatatypeAux.dtyp list -> Term.term list *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2175 | fun constructor_terms terms [] = terms | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2176 | | constructor_terms terms (d::ds) = | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2177 | let | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2178 | val dT = typ_of_dtyp descr typ_assoc d | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2179 | val d_terms = canonical_terms typs' dT | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2180 | in | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2181 | (* C_i x_1 ... x_n < C_i y_1 ... y_n if *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2182 | (* (x_1, ..., x_n) < (y_1, ..., y_n) *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2183 | constructor_terms | 
| 25538 | 2184 | (map_product (curry op $) terms d_terms) ds | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2185 | end | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2186 | in | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2187 | (* C_i ... < C_j ... if i < j *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2188 | List.concat (map (fn (cname, ctyps) => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2189 | let | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2190 | val cTerm = Const (cname, | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2191 | map (typ_of_dtyp descr typ_assoc) ctyps ---> T) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2192 | in | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2193 | constructor_terms [cTerm] ctyps | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2194 | end) constrs) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2195 | end) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2196 | | NONE => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2197 | (* not an inductive datatype; in this case the argument types in *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2198 | (* 'Ts' may not be IDTs either, so 'print' should be safe *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2199 | map (fn intr => print thy (typs, []) T intr (K false)) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2200 | (make_constants thy (typs, []) T)) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2201 | | _ => (* TFree ..., TVar ... *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2202 | map (fn intr => print thy (typs, []) T intr (K false)) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2203 | (make_constants thy (typs, []) T)) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2204 | val (typs, terms) = model | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2205 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2206 | case AList.lookup (op =) terms t of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2207 | SOME intr => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2208 | (* return an existing interpretation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2209 | SOME (intr, model, args) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2210 | | NONE => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2211 | (case t of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2212 | Const (s, T) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2213 | (case body_type T of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2214 | Type (s', Ts') => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2215 | (case DatatypePackage.get_datatype thy s' of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2216 | SOME info => (* body type is an inductive datatype *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2217 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2218 | val index = #index info | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2219 | val descr = #descr info | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2220 | val (_, dtyps, constrs) = lookup descr index | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2221 | val typ_assoc = dtyps ~~ Ts' | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2222 | (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2223 | val _ = if Library.exists (fn d => | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2224 | case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2225 | then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2226 |                   raise REFUTE ("IDT_constructor_interpreter",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2227 | "datatype argument (for type " | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2228 | ^ Sign.string_of_typ thy (Type (s', Ts')) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2229 | ^ ") is not a variable") | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2230 | else () | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2231 | (* split the constructors into those occuring before/after *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2232 | (* 'Const (s, T)' *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2233 | val (constrs1, constrs2) = take_prefix (fn (cname, ctypes) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2234 | not (cname = s andalso Sign.typ_instance thy (T, | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2235 | map (typ_of_dtyp descr typ_assoc) ctypes | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2236 | ---> Type (s', Ts')))) constrs | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2237 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2238 | case constrs2 of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2239 | [] => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2240 | (* 'Const (s, T)' is not a constructor of this datatype *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2241 | NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2242 | | (_, ctypes)::cs => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2243 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2244 | (* int option -- only /recursive/ IDTs have an associated *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2245 | (* depth *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2246 | val depth = AList.lookup (op =) typs (Type (s', Ts')) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2247 | (* this should never happen: at depth 0, this IDT fragment *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2248 | (* is definitely empty, and in this case we don't need to *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2249 | (* interpret its constructors *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2250 | val _ = (case depth of SOME 0 => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2251 |                       raise REFUTE ("IDT_constructor_interpreter",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2252 | "depth is 0") | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2253 | | _ => ()) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2254 | val typs' = (case depth of NONE => typs | SOME n => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2255 | AList.update (op =) (Type (s', Ts'), n-1) typs) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2256 | (* elements of the datatype come before elements generated *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2257 | (* by 'Const (s, T)' iff they are generated by a *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2258 | (* constructor in constrs1 *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2259 | val offset = size_of_dtyp thy typs' descr typ_assoc constrs1 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2260 | (* compute the total (current) size of the datatype *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2261 | val total = offset + | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2262 | size_of_dtyp thy typs' descr typ_assoc constrs2 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2263 | (* sanity check *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2264 | val _ = if total <> size_of_type thy (typs, []) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2265 | (Type (s', Ts')) then | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2266 |                       raise REFUTE ("IDT_constructor_interpreter",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2267 | "total is not equal to current size") | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2268 | else () | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2269 | (* returns an interpretation where everything is mapped to *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2270 | (* an "undefined" element of the datatype *) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2271 | (* DatatypeAux.dtyp list -> interpretation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2272 | fun make_undef [] = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2273 | Leaf (replicate total False) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2274 | | make_undef (d::ds) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2275 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2276 | (* compute the current size of the type 'd' *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2277 | val dT = typ_of_dtyp descr typ_assoc d | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2278 | val size = size_of_type thy (typs, []) dT | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2279 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2280 | Node (replicate size (make_undef ds)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2281 | end | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2282 | (* returns the interpretation for a constructor *) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2283 | (* int * DatatypeAux.dtyp list -> int * interpretation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2284 | fun make_constr (offset, []) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2285 | if offset<total then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2286 | (offset+1, Leaf ((replicate offset False) @ True :: | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2287 | (replicate (total-offset-1) False))) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2288 | else | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2289 |                       raise REFUTE ("IDT_constructor_interpreter",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2290 | "offset >= total") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2291 | | make_constr (offset, d::ds) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2292 | let | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2293 | (* Term.typ *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2294 | val dT = typ_of_dtyp descr typ_assoc d | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2295 | (* compute canonical term representations for all *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2296 | (* elements of the type 'd' (with the reduced depth *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2297 | (* for the IDT) *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2298 | val terms' = canonical_terms typs' dT | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2299 | (* sanity check *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2300 | val _ = if length terms' <> | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2301 | size_of_type thy (typs', []) dT | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2302 | then | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2303 |                           raise REFUTE ("IDT_constructor_interpreter",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2304 | "length of terms' is not equal to old size") | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2305 | else () | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2306 | (* compute canonical term representations for all *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2307 | (* elements of the type 'd' (with the current depth *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2308 | (* for the IDT) *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2309 | val terms = canonical_terms typs dT | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2310 | (* sanity check *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2311 | val _ = if length terms <> size_of_type thy (typs, []) dT | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2312 | then | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2313 |                           raise REFUTE ("IDT_constructor_interpreter",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2314 | "length of terms is not equal to current size") | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2315 | else () | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2316 | (* sanity check *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2317 | val _ = if length terms < length terms' then | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2318 |                           raise REFUTE ("IDT_constructor_interpreter",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2319 | "current size is less than old size") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2320 | else () | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2321 | (* sanity check: every element of terms' must also be *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2322 | (* present in terms *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2323 | val _ = if List.all (member op= terms) terms' then () | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2324 | else | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2325 |                           raise REFUTE ("IDT_constructor_interpreter",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2326 | "element has disappeared") | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2327 | (* sanity check: the order on elements of terms' is *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2328 | (* the same in terms, for those elements *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2329 | val _ = let | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2330 | fun search (x::xs) (y::ys) = | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2331 | if x = y then search xs ys else search (x::xs) ys | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2332 | | search (x::xs) [] = | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2333 |                             raise REFUTE ("IDT_constructor_interpreter",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2334 | "element order not preserved") | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2335 | | search [] _ = () | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2336 | in search terms' terms end | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2337 | (* int * interpretation list *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2338 | val (new_offset, intrs) = foldl_map (fn (off, t_elem) => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2339 | (* if 't_elem' existed at the previous depth, *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2340 | (* proceed recursively, otherwise map the entire *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2341 | (* subtree to "undefined" *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2342 | if t_elem mem terms' then | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2343 | make_constr (off, ds) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2344 | else | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2345 | (off, make_undef ds)) (offset, terms) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2346 | in | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2347 | (new_offset, Node intrs) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2348 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2349 | in | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2350 | SOME (snd (make_constr (offset, ctypes)), model, args) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2351 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2352 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2353 | | NONE => (* body type is not an inductive datatype *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2354 | NONE) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2355 | | _ => (* body type is a (free or schematic) type variable *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2356 | NONE) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2357 | | _ => (* term is not a constant *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2358 | NONE) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2359 | 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: 
14604diff
changeset | 2360 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2361 | (* theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2362 | (interpretation * model * arguments) option *) | 
| 14807 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 2363 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2364 | (* Difficult code ahead. Make sure you understand the *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2365 | (* 'IDT_constructor_interpreter' and the order in which it enumerates *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2366 | (* 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: 
15531diff
changeset | 2367 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2368 | fun IDT_recursion_interpreter thy model args t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2369 | (* careful: here we descend arbitrarily deep into 't', possibly before *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2370 | (* any other interpreter for atomic terms has had a chance to look at *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2371 | (* 't' *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2372 | case strip_comb t of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2373 | (Const (s, T), params) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2374 | (* iterate over all datatypes in 'thy' *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2375 | Symtab.fold (fn (_, info) => fn result => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2376 | case result of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2377 | SOME _ => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2378 | result (* just keep 'result' *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2379 | | NONE => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2380 | if member (op =) (#rec_names info) s then | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2381 | (* we do have a recursion operator of one of the (mutually *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2382 | (* recursive) datatypes given by 'info' *) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2383 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2384 | (* number of all constructors, including those of different *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2385 | (* (mutually recursive) datatypes within the same descriptor *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2386 | val mconstrs_count = sum (map (fn (_, (_, _, cs)) => length cs) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2387 | (#descr info)) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2388 | in | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2389 | if mconstrs_count < length params then | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2390 | (* too many actual parameters; for now we'll use the *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2391 | (* 'stlc_interpreter' to strip off one application *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2392 | NONE | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2393 | else if mconstrs_count > length params then | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2394 | (* too few actual parameters; we use eta expansion *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2395 | (* Note that the resulting expansion of lambda abstractions *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2396 | (* by the 'stlc_interpreter' may be rather slow (depending *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2397 | (* on the argument types and the size of the IDT, of *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2398 | (* course). *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2399 | SOME (interpret thy model args (eta_expand t | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2400 | (mconstrs_count - length params))) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2401 | else (* mconstrs_count = length params *) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2402 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2403 | (* interpret each parameter separately *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2404 | val ((model', args'), p_intrs) = foldl_map (fn ((m, a), p) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2405 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2406 | val (i, m', a') = interpret thy m a p | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2407 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2408 | ((m', a'), i) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2409 | end) ((model, args), params) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2410 | val (typs, _) = model' | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2411 | (* 'index' is /not/ necessarily the index of the IDT that *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2412 | (* the recursion operator is associated with, but merely *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2413 | (* the index of some mutually recursive IDT *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2414 | val index = #index info | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2415 | val descr = #descr info | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2416 | val (_, dtyps, _) = lookup descr index | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2417 | (* sanity check: we assume that the order of constructors *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2418 | (* in 'descr' is the same as the order of *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2419 | (* corresponding parameters, otherwise the *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2420 | (* association code below won't match the *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2421 | (* right constructors/parameters; we also *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2422 | (* assume that the order of recursion *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2423 | (* operators in '#rec_names info' is the *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2424 | (* same as the order of corresponding *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2425 | (* datatypes in 'descr' *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2426 | val _ = if map fst descr <> (0 upto (length descr - 1)) then | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2427 |                       raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2428 | "order of constructors and corresponding parameters/" ^ | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2429 | "recursion operators and corresponding datatypes " ^ | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2430 | "different?") | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2431 | else () | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2432 | (* sanity check: every element in 'dtyps' must be a *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2433 | (* 'DtTFree' *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2434 | val _ = if Library.exists (fn d => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2435 | case d of DatatypeAux.DtTFree _ => false | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2436 | | _ => true) dtyps | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2437 | then | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2438 |                       raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2439 | "datatype argument is not a variable") | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2440 | else () | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2441 | (* the type of a recursion operator is *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2442 | (* [T1, ..., Tn, IDT] ---> Tresult *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2443 | val IDT = List.nth (binder_types T, mconstrs_count) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2444 | (* by our assumption on the order of recursion operators *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2445 | (* and datatypes, this is the index of the datatype *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2446 | (* corresponding to the given recursion operator *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2447 | val idt_index = find_index_eq s (#rec_names info) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2448 | (* mutually recursive types must have the same type *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2449 | (* parameters, unless the mutual recursion comes from *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2450 | (* indirect recursion *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2451 | (* (DatatypeAux.dtyp * Term.typ) list -> | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2452 | (DatatypeAux.dtyp * Term.typ) list -> | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2453 | (DatatypeAux.dtyp * Term.typ) list *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2454 | fun rec_typ_assoc acc [] = | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2455 | acc | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2456 | | rec_typ_assoc acc ((d, T)::xs) = | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2457 | (case AList.lookup op= acc d of | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2458 | NONE => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2459 | (case d of | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2460 | DatatypeAux.DtTFree _ => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2461 | (* add the association, proceed *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2462 | rec_typ_assoc ((d, T)::acc) xs | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2463 | | DatatypeAux.DtType (s, ds) => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2464 | let | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2465 | val (s', Ts) = dest_Type T | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2466 | in | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2467 | if s=s' then | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2468 | rec_typ_assoc ((d, T)::acc) ((ds ~~ Ts) @ xs) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2469 | else | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2470 |                             raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2471 | "DtType/Type mismatch") | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2472 | end | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2473 | | DatatypeAux.DtRec i => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2474 | let | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2475 | val (_, ds, _) = lookup descr i | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2476 | val (_, Ts) = dest_Type T | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2477 | in | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2478 | rec_typ_assoc ((d, T)::acc) ((ds ~~ Ts) @ xs) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2479 | end) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2480 | | SOME T' => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2481 | if T=T' then | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2482 | (* ignore the association since it's already *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2483 | (* present, proceed *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2484 | rec_typ_assoc acc xs | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2485 | else | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2486 |                         raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2487 | "different type associations for the same dtyp")) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2488 | (* (DatatypeAux.dtyp * Term.typ) list *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2489 | val typ_assoc = List.filter | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2490 | (fn (DatatypeAux.DtTFree _, _) => true | (_, _) => false) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2491 | (rec_typ_assoc [] | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2492 | (#2 (lookup descr idt_index) ~~ (snd o dest_Type) IDT)) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2493 | (* sanity check: typ_assoc must associate types to the *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2494 | (* elements of 'dtyps' (and only to those) *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2495 | val _ = if not (Library.eq_set (dtyps, map fst typ_assoc)) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2496 | then | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2497 |                       raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2498 | "type association has extra/missing elements") | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2499 | else () | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2500 | (* interpret each constructor in the descriptor (including *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2501 | (* those of mutually recursive datatypes) *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2502 | (* (int * interpretation list) list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2503 | val mc_intrs = map (fn (idx, (_, _, cs)) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2504 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2505 | val c_return_typ = typ_of_dtyp descr typ_assoc | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2506 | (DatatypeAux.DtRec idx) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2507 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2508 | (idx, map (fn (cname, cargs) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2509 |                         (#1 o interpret thy (typs, []) {maxvars=0,
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2510 | def_eq=false, next_idx=1, bounds=[], | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2511 | wellformed=True}) (Const (cname, map (typ_of_dtyp | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2512 | descr typ_assoc) cargs ---> c_return_typ))) cs) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2513 | end) descr | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2514 | (* associate constructors with corresponding parameters *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2515 | (* (int * (interpretation * interpretation) list) list *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2516 | val (p_intrs', mc_p_intrs) = foldl_map | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2517 | (fn (p_intrs', (idx, c_intrs)) => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2518 | let | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2519 | val len = length c_intrs | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2520 | in | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2521 | (List.drop (p_intrs', len), | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2522 | (idx, c_intrs ~~ List.take (p_intrs', len))) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2523 | end) (p_intrs, mc_intrs) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2524 | (* sanity check: no 'p_intr' may be left afterwards *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2525 | val _ = if p_intrs' <> [] then | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2526 |                       raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2527 | "more parameter than constructor interpretations") | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2528 | else () | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2529 | (* The recursion operator, applied to 'mconstrs_count' *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2530 | (* arguments, is a function that maps every element of the *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2531 | (* inductive datatype to an element of some result type. *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2532 | (* Recursion operators for mutually recursive IDTs are *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2533 | (* translated simultaneously. *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2534 | (* Since the order on datatype elements is given by an *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2535 | (* order on constructors (and then by the order on *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2536 | (* argument tuples), we can simply copy corresponding *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2537 | (* subtrees from 'p_intrs', in the order in which they are *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2538 | (* given. *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2539 | (* interpretation * interpretation -> interpretation list *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2540 | fun ci_pi (Leaf xs, pi) = | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2541 | (* if the constructor does not match the arguments to a *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2542 | (* defined element of the IDT, the corresponding value *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2543 | (* of the parameter must be ignored *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2544 | if List.exists (equal True) xs then [pi] else [] | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2545 | | ci_pi (Node xs, Node ys) = | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2546 | List.concat (map ci_pi (xs ~~ ys)) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2547 | | ci_pi (Node _, Leaf _) = | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2548 |                     raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2549 | "constructor takes more arguments than the " ^ | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2550 | "associated parameter") | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2551 | (* (int * interpretation list) list *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2552 | val rec_operators = map (fn (idx, c_p_intrs) => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2553 | (idx, List.concat (map ci_pi c_p_intrs))) mc_p_intrs | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2554 | (* sanity check: every recursion operator must provide as *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2555 | (* many values as the corresponding datatype *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2556 | (* has elements *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2557 | val _ = map (fn (idx, intrs) => | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2558 | let | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2559 | val T = typ_of_dtyp descr typ_assoc | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2560 | (DatatypeAux.DtRec idx) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2561 | in | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2562 | if length intrs <> size_of_type thy (typs, []) T then | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2563 |                         raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2564 | "wrong number of interpretations for rec. operator") | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2565 | else () | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2566 | end) rec_operators | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2567 | (* For non-recursive datatypes, we are pretty much done at *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2568 | (* this point. For recursive datatypes however, we still *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2569 | (* need to apply the interpretations in 'rec_operators' to *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2570 | (* (recursively obtained) interpretations for recursive *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2571 | (* constructor arguments. To do so more efficiently, we *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2572 | (* copy 'rec_operators' into arrays first. Each Boolean *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2573 | (* indicates whether the recursive arguments have been *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2574 | (* considered already. *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2575 | (* (int * (bool * interpretation) Array.array) list *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2576 | val REC_OPERATORS = map (fn (idx, intrs) => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2577 | (idx, Array.fromList (map (pair false) intrs))) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2578 | rec_operators | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2579 | (* takes an interpretation, and if some leaf of this *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2580 | (* interpretation is the 'elem'-th element of the type, *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2581 | (* the indices of the arguments leading to this leaf are *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2582 | (* returned *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2583 | (* interpretation -> int -> int list option *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2584 | fun get_args (Leaf xs) elem = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2585 | if find_index_eq True xs = elem then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2586 | SOME [] | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2587 | else | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2588 | NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2589 | | get_args (Node xs) elem = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2590 | let | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2591 | (* interpretation list * int -> int list option *) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2592 | fun search ([], _) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2593 | NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2594 | | search (x::xs, n) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2595 | (case get_args x elem of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2596 | SOME result => SOME (n::result) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2597 | | NONE => search (xs, n+1)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2598 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2599 | search (xs, 0) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2600 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2601 | (* returns the index of the constructor and indices for *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2602 | (* its arguments that generate the 'elem'-th element of *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2603 | (* the datatype given by 'idx' *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2604 | (* int -> int -> int * int list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2605 | fun get_cargs idx elem = | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2606 | let | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2607 | (* int * interpretation list -> int * int list *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2608 | fun get_cargs_rec (_, []) = | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2609 |                       raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2610 | "no matching constructor found for datatype element") | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2611 | | get_cargs_rec (n, x::xs) = | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2612 | (case get_args x elem of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2613 | SOME args => (n, args) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2614 | | NONE => get_cargs_rec (n+1, xs)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2615 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2616 | get_cargs_rec (0, lookup mc_intrs idx) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2617 | end | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2618 | (* computes one entry in 'REC_OPERATORS', and recursively *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2619 | (* all entries needed for it, where 'idx' gives the *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2620 | (* datatype and 'elem' the element of it *) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2621 | (* int -> int -> interpretation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2622 | fun compute_array_entry idx elem = | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2623 | let | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2624 | val arr = lookup REC_OPERATORS idx | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2625 | val (flag, intr) = Array.sub (arr, elem) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2626 | in | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2627 | if flag then | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2628 | (* simply return the previously computed result *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2629 | intr | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2630 | else | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2631 | (* we have to apply 'intr' to interpretations for all *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2632 | (* recursive arguments *) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2633 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2634 | (* int * int list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2635 | val (c, args) = get_cargs idx elem | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2636 | (* find the indices of the constructor's /recursive/ *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2637 | (* arguments *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2638 | val (_, _, constrs) = lookup descr idx | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2639 | val (_, dtyps) = List.nth (constrs, c) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2640 | val rec_dtyps_args = List.filter | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2641 | (DatatypeAux.is_rec_type o fst) (dtyps ~~ args) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2642 | (* map those indices to interpretations *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2643 | (* (DatatypeAux.dtyp * interpretation) list *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2644 | val rec_dtyps_intrs = map (fn (dtyp, arg) => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2645 | let | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2646 | val dT = typ_of_dtyp descr typ_assoc dtyp | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2647 | val consts = make_constants thy (typs, []) dT | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2648 | val arg_i = List.nth (consts, arg) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2649 | in | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2650 | (dtyp, arg_i) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2651 | end) rec_dtyps_args | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2652 | (* takes the dtyp and interpretation of an element, *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2653 | (* and computes the interpretation for the *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2654 | (* corresponding recursive argument *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2655 | (* DatatypeAux.dtyp -> interpretation -> | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2656 | interpretation *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2657 | fun rec_intr (DatatypeAux.DtRec i) (Leaf xs) = | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2658 | (* recursive argument is "rec_i params elem" *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2659 | compute_array_entry i (find_index_eq True xs) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2660 | | rec_intr (DatatypeAux.DtRec _) (Node _) = | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2661 |                           raise REFUTE ("IDT_recursion_interpreter",
 | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2662 | "interpretation for IDT is a node") | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2663 |                           | rec_intr (DatatypeAux.DtType ("fun", [dt1, dt2]))
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2664 | (Node xs) = | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2665 | (* recursive argument is something like *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2666 | (* "\<lambda>x::dt1. rec_? params (elem x)" *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2667 | Node (map (rec_intr dt2) xs) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2668 |                           | rec_intr (DatatypeAux.DtType ("fun", [_, _]))
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2669 | (Leaf _) = | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2670 |                           raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2671 | "interpretation for function dtyp is a leaf") | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2672 | | rec_intr _ _ = | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2673 | (* admissibility ensures that every recursive type *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2674 | (* is of the form 'Dt_1 -> ... -> Dt_k -> *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2675 | (* (DtRec i)' *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2676 |                           raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2677 | "non-recursive codomain in recursive dtyp") | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2678 | (* obtain interpretations for recursive arguments *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2679 | (* interpretation list *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2680 | val arg_intrs = map (uncurry rec_intr) rec_dtyps_intrs | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2681 | (* apply 'intr' to all recursive arguments *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2682 | val result = foldl (fn (arg_i, i) => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2683 | interpretation_apply (i, arg_i)) intr arg_intrs | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2684 | (* update 'REC_OPERATORS' *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2685 | val _ = Array.update (arr, elem, (true, result)) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2686 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2687 | result | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2688 | end | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2689 | end | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2690 | val idt_size = Array.length (lookup REC_OPERATORS idt_index) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2691 | (* sanity check: the size of 'IDT' should be 'idt_size' *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2692 | val _ = if idt_size <> size_of_type thy (typs, []) IDT then | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2693 |                         raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2694 | "unexpected size of IDT (wrong type associated?)") | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2695 | else () | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2696 | (* interpretation *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2697 | val rec_op = Node (map (compute_array_entry idt_index) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2698 | (0 upto (idt_size - 1))) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2699 | in | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2700 | SOME (rec_op, model', args') | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2701 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2702 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2703 | else | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2704 | NONE (* not a recursion operator of this datatype *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2705 | ) (DatatypePackage.get_datatypes thy) NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2706 | | _ => (* head of term is not a constant *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2707 | 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: 
15531diff
changeset | 2708 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2709 | (* theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2710 | (interpretation * model * arguments) option *) | 
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15531diff
changeset | 2711 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2712 | (* only an optimization: 'card' could in principle be interpreted with *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2713 | (* interpreters available already (using its definition), but the code *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2714 | (* 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: 
14604diff
changeset | 2715 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2716 | fun Finite_Set_card_interpreter thy model args t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2717 | case t of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2718 |       Const ("Finite_Set.card",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2719 |         Type ("fun", [Type ("set", [T]), Type ("nat", [])])) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2720 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2721 | (* interpretation -> int *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2722 | fun number_of_elements (Node xs) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2723 | Library.foldl (fn (n, x) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2724 | if x=TT then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2725 | n+1 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2726 | else if x=FF then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2727 | n | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2728 | else | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2729 |               raise REFUTE ("Finite_Set_card_interpreter",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2730 | "interpretation for set type does not yield a Boolean")) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2731 | (0, xs) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2732 | | number_of_elements (Leaf _) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2733 |           raise REFUTE ("Finite_Set_card_interpreter",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2734 | "interpretation for set type is a leaf") | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2735 |         val size_of_nat = size_of_type thy model (Type ("nat", []))
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2736 | (* takes an interpretation for a set and returns an interpretation *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2737 | (* 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: 
22093diff
changeset | 2738 | (* interpretation -> interpretation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2739 | fun card i = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2740 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2741 | val n = number_of_elements i | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2742 | in | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2743 | if n<size_of_nat then | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2744 | Leaf ((replicate n False) @ True :: | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2745 | (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: 
22093diff
changeset | 2746 | else | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2747 | Leaf (replicate size_of_nat False) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2748 | end | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2749 |         val set_constants = make_constants thy model (Type ("set", [T]))
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2750 | in | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2751 | 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: 
22093diff
changeset | 2752 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2753 | | _ => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2754 | 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: 
14604diff
changeset | 2755 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2756 | (* theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2757 | (interpretation * model * arguments) option *) | 
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15531diff
changeset | 2758 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2759 | (* only an optimization: 'Finites' could in principle be interpreted with *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2760 | (* interpreters available already (using its definition), but the code *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2761 | (* below is more efficient *) | 
| 15571 | 2762 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2763 | fun Finite_Set_Finites_interpreter thy model args t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2764 | case t of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2765 |       Const ("Finite_Set.Finites", Type ("set", [Type ("set", [T])])) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2766 | let | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2767 |         val size_of_set = size_of_type thy model (Type ("set", [T]))
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2768 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2769 | (* we only consider finite models anyway, hence EVERY set is in *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2770 | (* "Finites" *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2771 | 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: 
22093diff
changeset | 2772 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2773 | | _ => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2774 | NONE; | 
| 15571 | 2775 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2776 | (* theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2777 | (interpretation * model * arguments) option *) | 
| 15571 | 2778 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2779 | (* only an optimization: 'finite' could in principle be interpreted with *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2780 | (* interpreters available already (using its definition), but the code *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2781 | (* below is more efficient *) | 
| 22093 | 2782 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2783 | fun Finite_Set_finite_interpreter thy model args t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2784 | case t of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2785 |       Const ("Finite_Set.finite",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2786 |         Type ("fun", [Type ("set", [T]), Type ("bool", [])])) $ _ =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2787 | (* 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: 
22093diff
changeset | 2788 | (* "finite" *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2789 | SOME (TT, model, args) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2790 |     | Const ("Finite_Set.finite",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2791 |         Type ("fun", [Type ("set", [T]), Type ("bool", [])])) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2792 | let | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2793 |         val size_of_set = size_of_type thy model (Type ("set", [T]))
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2794 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2795 | (* 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: 
22093diff
changeset | 2796 | (* "finite" *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2797 | 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: 
22093diff
changeset | 2798 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2799 | | _ => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2800 | NONE; | 
| 22093 | 2801 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2802 | (* theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2803 | (interpretation * model * arguments) option *) | 
| 22093 | 2804 | |
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2805 | (* only an optimization: 'HOL.less' could in principle be interpreted with *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2806 | (* interpreters available already (using its definition), but the code *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2807 | (* 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: 
15531diff
changeset | 2808 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2809 | fun Nat_less_interpreter thy model args t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2810 | case t of | 
| 23881 | 2811 |       Const (@{const_name HOL.less}, Type ("fun", [Type ("nat", []),
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2812 |         Type ("fun", [Type ("nat", []), Type ("bool", [])])])) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2813 | let | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2814 |         val size_of_nat = size_of_type thy model (Type ("nat", []))
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2815 | (* 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: 
24928diff
changeset | 2816 | (* is less than the remaining 'size_of_nat - n' nats *) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2817 | (* int -> interpretation *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2818 | 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: 
22093diff
changeset | 2819 | in | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2820 | 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: 
22093diff
changeset | 2821 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2822 | | _ => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2823 | 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: 
15531diff
changeset | 2824 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2825 | (* theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2826 | (interpretation * model * arguments) option *) | 
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15531diff
changeset | 2827 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2828 | (* only an optimization: 'HOL.plus' could in principle be interpreted with *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2829 | (* interpreters available already (using its definition), but the code *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2830 | (* 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: 
15531diff
changeset | 2831 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2832 | fun Nat_plus_interpreter thy model args t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2833 | case t of | 
| 22997 | 2834 |       Const (@{const_name HOL.plus}, Type ("fun", [Type ("nat", []),
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2835 |         Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2836 | let | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2837 |         val size_of_nat = size_of_type thy model (Type ("nat", []))
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2838 | (* int -> int -> interpretation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2839 | fun plus m n = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2840 | let | 
| 25032 | 2841 | val element = m + n | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2842 | in | 
| 25032 | 2843 | if element > size_of_nat - 1 then | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2844 | Leaf (replicate size_of_nat False) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2845 | else | 
| 25032 | 2846 | Leaf ((replicate element False) @ True :: | 
| 2847 | (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: 
22093diff
changeset | 2848 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2849 | in | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2850 | SOME (Node (map (fn m => Node (map (plus m) (0 upto size_of_nat-1))) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2851 | (0 upto size_of_nat-1)), model, args) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2852 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2853 | | _ => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2854 | 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: 
15531diff
changeset | 2855 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2856 | (* theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2857 | (interpretation * model * arguments) option *) | 
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15531diff
changeset | 2858 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2859 | (* only an optimization: 'HOL.minus' could in principle be interpreted *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2860 | (* with interpreters available already (using its definition), but the *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2861 | (* 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: 
15531diff
changeset | 2862 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2863 | fun Nat_minus_interpreter thy model args t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2864 | case t of | 
| 22997 | 2865 |       Const (@{const_name HOL.minus}, Type ("fun", [Type ("nat", []),
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2866 |         Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2867 | let | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2868 |         val size_of_nat = size_of_type thy model (Type ("nat", []))
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2869 | (* int -> int -> interpretation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2870 | fun minus m n = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2871 | let | 
| 25032 | 2872 | 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: 
22093diff
changeset | 2873 | in | 
| 25032 | 2874 | Leaf ((replicate element False) @ True :: | 
| 2875 | (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: 
22093diff
changeset | 2876 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2877 | in | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2878 | SOME (Node (map (fn m => Node (map (minus m) (0 upto size_of_nat-1))) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2879 | (0 upto size_of_nat-1)), model, args) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2880 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2881 | | _ => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2882 | 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: 
15531diff
changeset | 2883 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2884 | (* theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2885 | (interpretation * model * arguments) option *) | 
| 15547 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 webertj parents: 
15531diff
changeset | 2886 | |
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2887 | (* only an optimization: 'HOL.times' could in principle be interpreted *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2888 | (* with interpreters available already (using its definition), but the *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2889 | (* 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: 
15531diff
changeset | 2890 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2891 | fun Nat_times_interpreter thy model args t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2892 | case t of | 
| 22997 | 2893 |       Const (@{const_name HOL.times}, Type ("fun", [Type ("nat", []),
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2894 |         Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2895 | let | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2896 |         val size_of_nat = size_of_type thy model (Type ("nat", []))
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2897 | (* nat -> nat -> interpretation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2898 | fun mult m n = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2899 | let | 
| 25032 | 2900 | val element = m * n | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2901 | in | 
| 25032 | 2902 | if element > size_of_nat - 1 then | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2903 | Leaf (replicate size_of_nat False) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2904 | else | 
| 25032 | 2905 | Leaf ((replicate element False) @ True :: | 
| 2906 | (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: 
22093diff
changeset | 2907 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2908 | in | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2909 | SOME (Node (map (fn m => Node (map (mult m) (0 upto size_of_nat-1))) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2910 | (0 upto size_of_nat-1)), model, args) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2911 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2912 | | _ => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2913 | 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: 
15531diff
changeset | 2914 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2915 | (* theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2916 | (interpretation * model * arguments) option *) | 
| 15767 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15611diff
changeset | 2917 | |
| 23029 | 2918 | (* only an optimization: 'append' could in principle be interpreted with *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2919 | (* interpreters available already (using its definition), but the code *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 2920 | (* below is more efficient *) | 
| 15767 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15611diff
changeset | 2921 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2922 | fun List_append_interpreter thy model args t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2923 | case t of | 
| 23029 | 2924 |       Const ("List.append", Type ("fun", [Type ("List.list", [T]), Type ("fun",
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2925 |         [Type ("List.list", [_]), Type ("List.list", [_])])])) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2926 | let | 
| 25032 | 2927 | val size_elem = size_of_type thy model T | 
| 2928 |         val size_list   = size_of_type thy model (Type ("List.list", [T]))
 | |
| 2929 | (* maximal length of lists; 0 if we only consider the empty list *) | |
| 2930 | val list_length = let | |
| 2931 | (* int -> int -> int -> int *) | |
| 2932 | fun list_length_acc len lists total = | |
| 2933 | if lists = total then | |
| 2934 | len | |
| 2935 | else if lists < total then | |
| 2936 | list_length_acc (len+1) (lists*size_elem) (total-lists) | |
| 2937 | else | |
| 2938 |                 raise REFUTE ("List_append_interpreter",
 | |
| 2939 | "size_list not equal to 1 + size_elem + ... + " ^ | |
| 2940 | "size_elem^len, for some len") | |
| 2941 | in | |
| 2942 | list_length_acc 0 1 size_list | |
| 2943 | end | |
| 2944 | val elements = 0 upto (size_list-1) | |
| 2945 | (* FIXME: there should be a nice formula, which computes the same as *) | |
| 2946 | (* the following, but without all this intermediate tree *) | |
| 2947 | (* length/offset stuff *) | |
| 2948 | (* associate each list with its length and offset in a complete tree *) | |
| 2949 | (* of width 'size_elem' and depth 'length_list' (with 'size_list' *) | |
| 2950 | (* nodes total) *) | |
| 2951 | (* (int * (int * int)) list *) | |
| 2952 | val (_, lenoff_lists) = foldl_map (fn ((offsets, off), elem) => | |
| 2953 | (* 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: 
22093diff
changeset | 2954 | let | 
| 25032 | 2955 | val len = length offsets | 
| 2956 | (* associate the given element with len/off *) | |
| 2957 | val assoc = (elem, (len, off)) | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2958 | in | 
| 25032 | 2959 | if len < list_length then | 
| 2960 | (* go to first child node *) | |
| 2961 | ((off :: offsets, off * size_elem), assoc) | |
| 2962 | else if off mod size_elem < size_elem - 1 then | |
| 2963 | (* go to next sibling node *) | |
| 2964 | ((offsets, off + 1), assoc) | |
| 2965 | else | |
| 2966 | (* go back up the stack until we find a level where we can go *) | |
| 2967 | (* to the next sibling node *) | |
| 2968 | let | |
| 2969 | val offsets' = Library.dropwhile | |
| 2970 | (fn off' => off' mod size_elem = size_elem - 1) offsets | |
| 2971 | in | |
| 2972 | case offsets' of | |
| 2973 | [] => | |
| 2974 | (* we're at the last node in the tree; the next value *) | |
| 2975 | (* won't be used anyway *) | |
| 2976 | (([], 0), assoc) | |
| 2977 | | off'::offs' => | |
| 2978 | (* go to next sibling node *) | |
| 2979 | ((offs', off' + 1), assoc) | |
| 2980 | end | |
| 2981 | end) (([], 0), elements) | |
| 2982 | (* we also need the reverse association (from length/offset to *) | |
| 2983 | (* index) *) | |
| 2984 | val lenoff'_lists = map Library.swap lenoff_lists | |
| 2985 | (* 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: 
22093diff
changeset | 2986 | (* nat -> nat -> interpretation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2987 | fun append m n = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2988 | let | 
| 25032 | 2989 | val (len_m, off_m) = lookup lenoff_lists m | 
| 2990 | val (len_n, off_n) = lookup lenoff_lists n | |
| 2991 | val len_elem = len_m + len_n | |
| 2992 | val off_elem = off_m * power (size_elem, len_n) + off_n | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2993 | in | 
| 25032 | 2994 | case AList.lookup op= lenoff'_lists (len_elem, off_elem) of | 
| 2995 | NONE => | |
| 2996 | (* undefined *) | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 2997 | Leaf (replicate size_list False) | 
| 25032 | 2998 | | SOME element => | 
| 2999 | Leaf ((replicate element False) @ True :: | |
| 3000 | (replicate (size_list - element - 1) False)) | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3001 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3002 | in | 
| 25032 | 3003 | SOME (Node (map (fn m => Node (map (append m) elements)) elements), | 
| 3004 | model, args) | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3005 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3006 | | _ => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3007 | NONE; | 
| 15767 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 webertj parents: 
15611diff
changeset | 3008 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3009 | (* theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3010 | (interpretation * model * arguments) option *) | 
| 16050 | 3011 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3012 | (* only an optimization: 'lfp' could in principle be interpreted with *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3013 | (* interpreters available already (using its definition), but the code *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3014 | (* below is more efficient *) | 
| 16050 | 3015 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3016 | fun Lfp_lfp_interpreter thy model args t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3017 | case t of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3018 |       Const ("Lfp.lfp", Type ("fun", [Type ("fun",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3019 |         [Type ("set", [T]), Type ("set", [_])]), Type ("set", [_])])) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3020 | let | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3021 | val size_elem = size_of_type thy model T | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3022 | (* the universe (i.e. the set that contains every element) *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3023 | 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: 
22093diff
changeset | 3024 | (* all sets with elements from type 'T' *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3025 |         val i_sets = make_constants thy model (Type ("set", [T]))
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3026 | (* all functions that map sets to sets *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3027 |         val i_funs = make_constants thy model (Type ("fun",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3028 |           [Type ("set", [T]), Type ("set", [T])]))
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3029 |         (* "lfp(f) == Inter({u. f(u) <= u})" *)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3030 | (* interpretation * interpretation -> bool *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3031 | fun is_subset (Node subs, Node sups) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3032 | List.all (fn (sub, sup) => (sub = FF) orelse (sup = TT)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3033 | (subs ~~ sups) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3034 | | is_subset (_, _) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3035 |           raise REFUTE ("Lfp_lfp_interpreter",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3036 | "is_subset: interpretation for set is not a node") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3037 | (* interpretation * interpretation -> interpretation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3038 | fun intersection (Node xs, Node ys) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3039 | Node (map (fn (x, y) => if x=TT andalso y=TT then TT else FF) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3040 | (xs ~~ ys)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3041 | | intersection (_, _) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3042 |           raise REFUTE ("Lfp_lfp_interpreter",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3043 | "intersection: interpretation for set is not a node") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3044 | (* interpretation -> interpretaion *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3045 | fun lfp (Node resultsets) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3046 | foldl (fn ((set, resultset), acc) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3047 | if is_subset (resultset, set) then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3048 | intersection (acc, set) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3049 | else | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3050 | acc) i_univ (i_sets ~~ resultsets) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3051 | | lfp _ = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3052 |             raise REFUTE ("Lfp_lfp_interpreter",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3053 | "lfp: interpretation for function is not a node") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3054 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3055 | 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: 
22093diff
changeset | 3056 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3057 | | _ => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3058 | NONE; | 
| 16050 | 3059 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3060 | (* theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3061 | (interpretation * model * arguments) option *) | 
| 16050 | 3062 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3063 | (* only an optimization: 'gfp' could in principle be interpreted with *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3064 | (* interpreters available already (using its definition), but the code *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3065 | (* below is more efficient *) | 
| 16050 | 3066 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3067 | fun Gfp_gfp_interpreter thy model args t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3068 | case t of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3069 |       Const ("Gfp.gfp", Type ("fun", [Type ("fun",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3070 |         [Type ("set", [T]), Type ("set", [_])]), Type ("set", [_])])) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3071 | let nonfix union (* because "union" is used below *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3072 | val size_elem = size_of_type thy model T | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3073 | (* the universe (i.e. the set that contains every element) *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3074 | 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: 
22093diff
changeset | 3075 | (* all sets with elements from type 'T' *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3076 |         val i_sets = make_constants thy model (Type ("set", [T]))
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3077 | (* all functions that map sets to sets *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3078 |         val i_funs = make_constants thy model (Type ("fun",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3079 |           [Type ("set", [T]), Type ("set", [T])]))
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3080 |         (* "gfp(f) == Union({u. u <= f(u)})" *)
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3081 | (* interpretation * interpretation -> bool *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3082 | fun is_subset (Node subs, Node sups) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3083 | List.all (fn (sub, sup) => (sub = FF) orelse (sup = TT)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3084 | (subs ~~ sups) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3085 | | is_subset (_, _) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3086 |           raise REFUTE ("Gfp_gfp_interpreter",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3087 | "is_subset: interpretation for set is not a node") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3088 | (* interpretation * interpretation -> interpretation *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3089 | fun union (Node xs, Node ys) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3090 | Node (map (fn (x,y) => if x=TT orelse y=TT then TT else FF) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3091 | (xs ~~ ys)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3092 | | union (_, _) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3093 |           raise REFUTE ("Gfp_gfp_interpreter",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3094 | "union: interpretation for set is not a node") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3095 | (* interpretation -> interpretaion *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3096 | fun gfp (Node resultsets) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3097 | foldl (fn ((set, resultset), acc) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3098 | if is_subset (set, resultset) then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3099 | union (acc, set) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3100 | else | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3101 | acc) i_univ (i_sets ~~ resultsets) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3102 | | gfp _ = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3103 |             raise REFUTE ("Gfp_gfp_interpreter",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3104 | "gfp: interpretation for function is not a node") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3105 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3106 | SOME (Node (map gfp i_funs), model, args) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3107 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3108 | | _ => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3109 | NONE; | 
| 16050 | 3110 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3111 | (* theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3112 | (interpretation * model * arguments) option *) | 
| 21267 | 3113 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3114 | (* only an optimization: 'fst' could in principle be interpreted with *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3115 | (* interpreters available already (using its definition), but the code *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3116 | (* below is more efficient *) | 
| 21267 | 3117 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3118 | fun Product_Type_fst_interpreter thy model args t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3119 | case t of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3120 |       Const ("fst", Type ("fun", [Type ("*", [T, U]), _])) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3121 | let | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3122 | val constants_T = make_constants thy model T | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3123 | val size_U = size_of_type thy model U | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3124 | in | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3125 | SOME (Node (List.concat (map (replicate size_U) constants_T)), | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3126 | model, args) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3127 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3128 | | _ => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3129 | NONE; | 
| 21267 | 3130 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3131 | (* theory -> model -> arguments -> Term.term -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3132 | (interpretation * model * arguments) option *) | 
| 21267 | 3133 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3134 | (* only an optimization: 'snd' could in principle be interpreted with *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3135 | (* interpreters available already (using its definition), but the code *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3136 | (* below is more efficient *) | 
| 21267 | 3137 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3138 | fun Product_Type_snd_interpreter thy model args t = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3139 | case t of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3140 |       Const ("snd", Type ("fun", [Type ("*", [T, U]), _])) =>
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3141 | let | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3142 | val size_T = size_of_type thy model T | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3143 | val constants_U = make_constants thy model U | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3144 | in | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3145 | SOME (Node (List.concat (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: 
22093diff
changeset | 3146 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3147 | | _ => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3148 | NONE; | 
| 21267 | 3149 | |
| 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: 
14604diff
changeset | 3150 | |
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 3151 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 3152 | (* 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: 
14604diff
changeset | 3153 | (* ------------------------------------------------------------------------- *) | 
| 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 3154 | |
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3155 | (* theory -> model -> Term.typ -> interpretation -> (int -> bool) -> | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3156 | Term.term option *) | 
| 14807 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 webertj parents: 
14604diff
changeset | 3157 | |
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3158 | fun stlc_printer thy model T intr assignment = | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3159 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3160 | (* string -> string *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3161 | fun strip_leading_quote s = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3162 | (implode o (fn [] => [] | x::xs => if x="'" then xs else x::xs) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3163 | o explode) s | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3164 | (* Term.typ -> string *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3165 | fun string_of_typ (Type (s, _)) = s | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3166 | | string_of_typ (TFree (x, _)) = strip_leading_quote x | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3167 | | string_of_typ (TVar ((x,i), _)) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3168 | strip_leading_quote x ^ string_of_int i | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3169 | (* interpretation -> int *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3170 | fun index_from_interpretation (Leaf xs) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3171 | find_index (PropLogic.eval assignment) xs | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3172 | | index_from_interpretation _ = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3173 |       raise REFUTE ("stlc_printer",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3174 | "interpretation for ground type is not a leaf") | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3175 | in | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3176 | case T of | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3177 |       Type ("fun", [T1, T2]) =>
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3178 | let | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3179 | (* create all constants of type 'T1' *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3180 | val constants = make_constants thy model T1 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3181 | (* interpretation list *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3182 | val results = (case intr of | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3183 | Node xs => xs | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3184 |           | _       => raise REFUTE ("stlc_printer",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3185 | "interpretation for function type is a leaf")) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3186 | (* Term.term list *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3187 | val pairs = map (fn (arg, result) => | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3188 | HOLogic.mk_prod | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3189 | (print thy model T1 arg assignment, | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3190 | print thy model T2 result assignment)) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3191 | (constants ~~ results) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3192 | (* Term.typ *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3193 | val HOLogic_prodT = HOLogic.mk_prodT (T1, T2) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3194 | val HOLogic_setT = HOLogic.mk_setT HOLogic_prodT | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3195 | (* Term.term *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3196 |         val HOLogic_empty_set = Const ("{}", HOLogic_setT)
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3197 | val HOLogic_insert = | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3198 |           Const ("insert", HOLogic_prodT --> HOLogic_setT --> HOLogic_setT)
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3199 | in | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3200 | SOME (foldr (fn (pair, acc) => HOLogic_insert $ pair $ acc) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3201 | HOLogic_empty_set pairs) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3202 | end | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3203 |     | Type ("prop", [])      =>
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3204 | (case index_from_interpretation intr of | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3205 |         ~1 => SOME (HOLogic.mk_Trueprop (Const ("arbitrary", HOLogic.boolT)))
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3206 | | 0 => SOME (HOLogic.mk_Trueprop HOLogic.true_const) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3207 | | 1 => SOME (HOLogic.mk_Trueprop HOLogic.false_const) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3208 |       | _  => raise REFUTE ("stlc_interpreter",
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3209 | "illegal interpretation for a propositional value")) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3210 | | Type _ => if index_from_interpretation intr = (~1) then | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3211 |         SOME (Const ("arbitrary", T))
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3212 | else | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3213 | SOME (Const (string_of_typ T ^ | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3214 | string_of_int (index_from_interpretation intr), T)) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3215 | | TFree _ => if index_from_interpretation intr = (~1) then | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3216 |         SOME (Const ("arbitrary", T))
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3217 | else | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3218 | SOME (Const (string_of_typ T ^ | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3219 | string_of_int (index_from_interpretation intr), T)) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3220 | | TVar _ => if index_from_interpretation intr = (~1) then | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3221 |         SOME (Const ("arbitrary", T))
 | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3222 | else | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3223 | SOME (Const (string_of_typ T ^ | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3224 | 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: 
22093diff
changeset | 3225 | end; | 
| 14350 | 3226 | |
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3227 | (* theory -> model -> Term.typ -> interpretation -> (int -> bool) -> | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3228 | Term.term option *) | 
| 14350 | 3229 | |
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3230 | fun set_printer thy model T intr assignment = | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3231 | (case T of | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3232 |       Type ("set", [T1]) =>
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3233 | let | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3234 | (* create all constants of type 'T1' *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3235 | val constants = make_constants thy model T1 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3236 | (* interpretation list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3237 | val results = (case intr of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3238 | Node xs => xs | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3239 |           | _       => raise REFUTE ("set_printer",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3240 | "interpretation for set type is a leaf")) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3241 | (* Term.term list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3242 | val elements = List.mapPartial (fn (arg, result) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3243 | case result of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3244 | Leaf [fmTrue, fmFalse] => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3245 | if PropLogic.eval assignment fmTrue then | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3246 | SOME (print thy model T1 arg assignment) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3247 | else (* if PropLogic.eval assignment fmFalse then *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3248 | NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3249 | | _ => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3250 |             raise REFUTE ("set_printer",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3251 | "illegal interpretation for a Boolean value")) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3252 | (constants ~~ results) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3253 | (* Term.typ *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3254 | val HOLogic_setT1 = HOLogic.mk_setT T1 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3255 | (* Term.term *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3256 |         val HOLogic_empty_set = Const ("{}", HOLogic_setT1)
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3257 | val HOLogic_insert = | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3258 |           Const ("insert", T1 --> HOLogic_setT1 --> HOLogic_setT1)
 | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3259 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3260 | SOME (Library.foldl (fn (acc, elem) => HOLogic_insert $ elem $ acc) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3261 | (HOLogic_empty_set, elements)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3262 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3263 | | _ => | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3264 | NONE); | 
| 14350 | 3265 | |
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3266 | (* theory -> model -> Term.typ -> interpretation -> (int -> bool) -> | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3267 | Term.term option *) | 
| 14350 | 3268 | |
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3269 | fun IDT_printer thy model T intr assignment = | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3270 | (case T of | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3271 | Type (s, Ts) => | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3272 | (case DatatypePackage.get_datatype thy s of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3273 | SOME info => (* inductive datatype *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3274 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3275 | val (typs, _) = model | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3276 | val index = #index info | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3277 | val descr = #descr info | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3278 | val (_, dtyps, constrs) = lookup descr index | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3279 | val typ_assoc = dtyps ~~ Ts | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3280 | (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3281 | val _ = if Library.exists (fn d => | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3282 | case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3283 | then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3284 |               raise REFUTE ("IDT_printer", "datatype argument (for type " ^
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3285 | Sign.string_of_typ thy (Type (s, Ts)) ^ ") is not a variable") | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3286 | else () | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3287 | (* the index of the element in the datatype *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3288 | val element = (case intr of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3289 | Leaf xs => find_index (PropLogic.eval assignment) xs | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3290 |             | Node _  => raise REFUTE ("IDT_printer",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3291 | "interpretation is not a leaf")) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3292 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3293 | if element < 0 then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3294 |             SOME (Const ("arbitrary", Type (s, Ts)))
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3295 | else let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3296 | (* takes a datatype constructor, and if for some arguments this *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3297 | (* constructor generates the datatype's element that is given by *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3298 | (* 'element', returns the constructor (as a term) as well as the *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3299 | (* indices of the arguments *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3300 | (* string * DatatypeAux.dtyp list -> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3301 | (Term.term * int list) option *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3302 | fun get_constr_args (cname, cargs) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3303 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3304 | val cTerm = Const (cname, | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3305 | map (typ_of_dtyp descr typ_assoc) cargs ---> Type (s, Ts)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3306 |                 val (iC, _, _) = interpret thy (typs, []) {maxvars=0,
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3307 | def_eq=false, next_idx=1, bounds=[], wellformed=True} cTerm | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3308 | (* interpretation -> int list option *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3309 | fun get_args (Leaf xs) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3310 | if find_index_eq True xs = element then | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3311 | SOME [] | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3312 | else | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3313 | NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3314 | | get_args (Node xs) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3315 | let | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3316 | (* interpretation * int -> int list option *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3317 | fun search ([], _) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3318 | NONE | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3319 | | search (x::xs, n) = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3320 | (case get_args x of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3321 | SOME result => SOME (n::result) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3322 | | NONE => search (xs, n+1)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3323 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3324 | search (xs, 0) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3325 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3326 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3327 | Option.map (fn args => (cTerm, cargs, args)) (get_args iC) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3328 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3329 | (* Term.term * DatatypeAux.dtyp list * int list *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3330 | val (cTerm, cargs, args) = | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3331 | (* we could speed things up by computing the correct *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3332 | (* constructor directly (rather than testing all *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3333 | (* constructors), based on the order in which constructors *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3334 | (* generate elements of datatypes; the current implementation *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3335 | (* of 'IDT_printer' however is independent of the internals *) | 
| 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3336 | (* of 'IDT_constructor_interpreter' *) | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3337 | (case get_first get_constr_args constrs of | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3338 | SOME x => x | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3339 |               | NONE   => raise REFUTE ("IDT_printer",
 | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3340 | "no matching constructor found for element " ^ | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3341 | string_of_int element)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3342 | val argsTerms = map (fn (d, n) => | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3343 | let | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3344 | val dT = typ_of_dtyp descr typ_assoc d | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3345 | (* we only need the n-th element of this list, so there *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3346 | (* might be a more efficient implementation that does not *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3347 | (* generate all constants *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3348 | val consts = make_constants thy (typs, []) dT | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3349 | in | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3350 | print thy (typs, []) dT (List.nth (consts, n)) assignment | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3351 | end) (cargs ~~ args) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3352 | in | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3353 | SOME (Library.foldl op$ (cTerm, argsTerms)) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3354 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3355 | end | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3356 | | NONE => (* not an inductive datatype *) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3357 | NONE) | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3358 | | _ => (* a (free or schematic) type variable *) | 
| 25014 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 webertj parents: 
24928diff
changeset | 3359 | NONE); | 
| 14350 | 3360 | |
| 3361 | ||
| 3362 | (* ------------------------------------------------------------------------- *) | |
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 3363 | (* 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: 
14351diff
changeset | 3364 | (* structure *) | 
| 14350 | 3365 | (* ------------------------------------------------------------------------- *) | 
| 3366 | ||
| 3367 | (* ------------------------------------------------------------------------- *) | |
| 14456 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 3368 | (* Note: the interpreters and printers are used in reverse order; however, *) | 
| 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 webertj parents: 
14351diff
changeset | 3369 | (* 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: 
14604diff
changeset | 3370 | (* 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: 
14604diff
changeset | 3371 | (* subterms that are then passed to other interpreters! *) | 
| 14350 | 3372 | (* ------------------------------------------------------------------------- *) | 
| 3373 | ||
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3374 | (* (theory -> theory) list *) | 
| 14350 | 3375 | |
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3376 | val setup = | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3377 | add_interpreter "stlc" stlc_interpreter #> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3378 | add_interpreter "Pure" Pure_interpreter #> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3379 | add_interpreter "HOLogic" HOLogic_interpreter #> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3380 | add_interpreter "set" set_interpreter #> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3381 | add_interpreter "IDT" IDT_interpreter #> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3382 | add_interpreter "IDT_constructor" IDT_constructor_interpreter #> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3383 | add_interpreter "IDT_recursion" IDT_recursion_interpreter #> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3384 | add_interpreter "Finite_Set.card" Finite_Set_card_interpreter #> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3385 | add_interpreter "Finite_Set.Finites" Finite_Set_Finites_interpreter #> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3386 | add_interpreter "Finite_Set.finite" Finite_Set_finite_interpreter #> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3387 | add_interpreter "Nat_Orderings.less" Nat_less_interpreter #> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3388 | add_interpreter "Nat_HOL.plus" Nat_plus_interpreter #> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3389 | add_interpreter "Nat_HOL.minus" Nat_minus_interpreter #> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3390 | add_interpreter "Nat_HOL.times" Nat_times_interpreter #> | 
| 25032 | 3391 | add_interpreter "List.append" List_append_interpreter #> | 
| 22567 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3392 | add_interpreter "Lfp.lfp" Lfp_lfp_interpreter #> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3393 | add_interpreter "Gfp.gfp" Gfp_gfp_interpreter #> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3394 | add_interpreter "fst" Product_Type_fst_interpreter #> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3395 | add_interpreter "snd" Product_Type_snd_interpreter #> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3396 | add_printer "stlc" stlc_printer #> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3397 | add_printer "set" set_printer #> | 
| 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 wenzelm parents: 
22093diff
changeset | 3398 | add_printer "IDT" IDT_printer; | 
| 14350 | 3399 | |
| 22092 | 3400 | end (* structure Refute *) |