| author | wenzelm | 
| Mon, 19 Jun 2017 20:32:06 +0200 | |
| changeset 66117 | e6f808d1307c | 
| parent 62519 | a564458f94db | 
| child 67399 | eab6ce8368fa | 
| permissions | -rw-r--r-- | 
| 50530 | 1  | 
(* Title: HOL/Library/refute.ML  | 
| 
29265
 
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
 
wenzelm 
parents: 
29004 
diff
changeset
 | 
2  | 
Author: Tjark Weber, TU Muenchen  | 
| 14350 | 3  | 
|
| 14965 | 4  | 
Finite model generation for HOL formulas, using a SAT solver.  | 
| 14350 | 5  | 
*)  | 
6  | 
||
7  | 
signature REFUTE =  | 
|
8  | 
sig  | 
|
9  | 
||
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
10  | 
exception REFUTE of string * string  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
11  | 
|
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
12  | 
(* ------------------------------------------------------------------------- *)  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
13  | 
(* Model/interpretation related code (translation HOL -> propositional logic *)  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
14  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
15  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
16  | 
type params  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
17  | 
type interpretation  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
18  | 
type model  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
19  | 
type arguments  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
20  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
21  | 
exception MAXVARS_EXCEEDED  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
22  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
23  | 
val add_interpreter : string -> (Proof.context -> model -> arguments -> term ->  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
24  | 
(interpretation * model * arguments) option) -> theory -> theory  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
25  | 
val add_printer : string -> (Proof.context -> model -> typ ->  | 
| 33243 | 26  | 
interpretation -> (int -> bool) -> term option) -> theory -> theory  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
27  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
28  | 
val interpret : Proof.context -> model -> arguments -> term ->  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
29  | 
(interpretation * model * arguments)  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
30  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
31  | 
val print : Proof.context -> model -> typ -> interpretation -> (int -> bool) -> term  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
32  | 
val print_model : Proof.context -> model -> (int -> bool) -> string  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
33  | 
|
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
34  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
35  | 
(* Interface *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
36  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
37  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
38  | 
val set_default_param : (string * string) -> theory -> theory  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
39  | 
val get_default_param : Proof.context -> string -> string option  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
40  | 
val get_default_params : Proof.context -> (string * string) list  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
41  | 
val actual_params : Proof.context -> (string * string) list -> params  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
42  | 
|
| 
45387
 
ccffb3f9f42b
return outcome code, so that it can be picked up by Mutabelle
 
blanchet 
parents: 
44121 
diff
changeset
 | 
43  | 
val find_model :  | 
| 
 
ccffb3f9f42b
return outcome code, so that it can be picked up by Mutabelle
 
blanchet 
parents: 
44121 
diff
changeset
 | 
44  | 
Proof.context -> params -> term list -> term -> bool -> string  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
45  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
46  | 
(* tries to find a model for a formula: *)  | 
| 
34120
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
47  | 
val satisfy_term :  | 
| 
45387
 
ccffb3f9f42b
return outcome code, so that it can be picked up by Mutabelle
 
blanchet 
parents: 
44121 
diff
changeset
 | 
48  | 
Proof.context -> (string * string) list -> term list -> term -> string  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
49  | 
(* tries to find a model that refutes a formula: *)  | 
| 
34120
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
50  | 
val refute_term :  | 
| 
45387
 
ccffb3f9f42b
return outcome code, so that it can be picked up by Mutabelle
 
blanchet 
parents: 
44121 
diff
changeset
 | 
51  | 
Proof.context -> (string * string) list -> term list -> term -> string  | 
| 
34120
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
52  | 
val refute_goal :  | 
| 
45387
 
ccffb3f9f42b
return outcome code, so that it can be picked up by Mutabelle
 
blanchet 
parents: 
44121 
diff
changeset
 | 
53  | 
Proof.context -> (string * string) list -> thm -> int -> string  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
54  | 
|
| 
29802
 
d201a838d0f7
Make some Refute functions public so I can use them in Nitpick,
 
blanchet 
parents: 
29288 
diff
changeset
 | 
55  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
d201a838d0f7
Make some Refute functions public so I can use them in Nitpick,
 
blanchet 
parents: 
29288 
diff
changeset
 | 
56  | 
(* Additional functions used by Nitpick (to be factored out) *)  | 
| 
 
d201a838d0f7
Make some Refute functions public so I can use them in Nitpick,
 
blanchet 
parents: 
29288 
diff
changeset
 | 
57  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
d201a838d0f7
Make some Refute functions public so I can use them in Nitpick,
 
blanchet 
parents: 
29288 
diff
changeset
 | 
58  | 
|
| 33243 | 59  | 
val get_classdef : theory -> string -> (string * term) option  | 
60  | 
val norm_rhs : term -> term  | 
|
61  | 
val get_def : theory -> string * typ -> (string * term) option  | 
|
62  | 
val get_typedef : theory -> typ -> (string * term) option  | 
|
63  | 
val is_IDT_constructor : theory -> string * typ -> bool  | 
|
64  | 
val is_IDT_recursor : theory -> string * typ -> bool  | 
|
65  | 
val is_const_of_class: theory -> string * typ -> bool  | 
|
66  | 
val string_of_typ : typ -> string  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
67  | 
end;  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
68  | 
|
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
69  | 
structure Refute : REFUTE =  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
70  | 
struct  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
71  | 
|
| 41471 | 72  | 
open Prop_Logic;  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
73  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
74  | 
(* We use 'REFUTE' only for internal error conditions that should *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
75  | 
(* never occur in the first place (i.e. errors caused by bugs in our *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
76  | 
(* code). Otherwise (e.g. to indicate invalid input data) we use *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
77  | 
(* 'error'. *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
78  | 
exception REFUTE of string * string;  (* ("in function", "cause") *)
 | 
| 14350 | 79  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
80  | 
(* should be raised by an interpreter when more variables would be *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
81  | 
(* required than allowed by 'maxvars' *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
82  | 
exception MAXVARS_EXCEEDED;  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
83  | 
|
| 14350 | 84  | 
|
85  | 
(* ------------------------------------------------------------------------- *)  | 
|
86  | 
(* TREES *)  | 
|
87  | 
(* ------------------------------------------------------------------------- *)  | 
|
88  | 
||
89  | 
(* ------------------------------------------------------------------------- *)  | 
|
90  | 
(* tree: implements an arbitrarily (but finitely) branching tree as a list *)  | 
|
91  | 
(* of (lists of ...) elements *)  | 
|
92  | 
(* ------------------------------------------------------------------------- *)  | 
|
93  | 
||
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
94  | 
datatype 'a tree =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
95  | 
Leaf of 'a  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
96  | 
  | Node of ('a tree) list;
 | 
| 14350 | 97  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
98  | 
fun tree_map f tr =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
99  | 
case tr of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
100  | 
Leaf x => Leaf (f x)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
101  | 
| Node xs => Node (map (tree_map f) xs);  | 
| 14350 | 102  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
103  | 
fun tree_pair (t1, t2) =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
104  | 
case t1 of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
105  | 
Leaf x =>  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
106  | 
(case t2 of  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
107  | 
Leaf y => Leaf (x,y)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
108  | 
        | Node _ => raise REFUTE ("tree_pair",
 | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
109  | 
"trees are of different height (second tree is higher)"))  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
110  | 
| Node xs =>  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
111  | 
(case t2 of  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
112  | 
(* '~~' will raise an exception if the number of branches in *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
113  | 
(* both trees is different at the current node *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
114  | 
Node ys => Node (map tree_pair (xs ~~ ys))  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
115  | 
        | Leaf _  => raise REFUTE ("tree_pair",
 | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
116  | 
"trees are of different height (first tree is higher)"));  | 
| 14350 | 117  | 
|
118  | 
(* ------------------------------------------------------------------------- *)  | 
|
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
119  | 
(* params: parameters that control the translation into a propositional *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
120  | 
(* formula/model generation *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
121  | 
(* *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
122  | 
(* The following parameters are supported (and required (!), except for *)  | 
| 
30314
 
853778f6ef7d
Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
 
blanchet 
parents: 
30312 
diff
changeset
 | 
123  | 
(* "sizes" and "expect"): *)  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
124  | 
(* *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
125  | 
(* Name Type Description *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
126  | 
(* *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
127  | 
(* "sizes" (string * int) list *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
128  | 
(* Size of ground types (e.g. 'a=2), or depth of IDTs. *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
129  | 
(* "minsize" int If >0, minimal size of each ground type/IDT depth. *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
130  | 
(* "maxsize" int If >0, maximal size of each ground type/IDT depth. *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
131  | 
(* "maxvars" int If >0, use at most 'maxvars' Boolean variables *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
132  | 
(* when transforming the term into a propositional *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
133  | 
(* formula. *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
134  | 
(* "maxtime" int If >0, terminate after at most 'maxtime' seconds. *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
135  | 
(* "satsolver" string SAT solver to be used. *)  | 
| 
34120
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
136  | 
(* "no_assms" bool If "true", assumptions in structured proofs are *)  | 
| 
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
137  | 
(* not considered. *)  | 
| 
30314
 
853778f6ef7d
Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
 
blanchet 
parents: 
30312 
diff
changeset
 | 
138  | 
(* "expect"      string  Expected result ("genuine", "potential", "none", or *)
 | 
| 
34120
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
139  | 
(* "unknown"). *)  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
140  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
141  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
142  | 
type params =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
143  | 
  {
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
144  | 
sizes : (string * int) list,  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
145  | 
minsize : int,  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
146  | 
maxsize : int,  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
147  | 
maxvars : int,  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
148  | 
maxtime : int,  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
149  | 
satsolver: string,  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
150  | 
no_assms : bool,  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
151  | 
expect : string  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
152  | 
};  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
153  | 
|
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
154  | 
(* ------------------------------------------------------------------------- *)  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
155  | 
(* interpretation: a term's interpretation is given by a variable of type *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
156  | 
(* 'interpretation' *)  | 
| 14350 | 157  | 
(* ------------------------------------------------------------------------- *)  | 
158  | 
||
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
159  | 
type interpretation =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
160  | 
prop_formula list tree;  | 
| 14350 | 161  | 
|
162  | 
(* ------------------------------------------------------------------------- *)  | 
|
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
163  | 
(* model: a model specifies the size of types and the interpretation of *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
164  | 
(* terms *)  | 
| 14350 | 165  | 
(* ------------------------------------------------------------------------- *)  | 
166  | 
||
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
167  | 
type model =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
168  | 
(typ * int) list * (term * interpretation) list;  | 
| 14350 | 169  | 
|
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
170  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
171  | 
(* arguments: additional arguments required during interpretation of terms *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
172  | 
(* ------------------------------------------------------------------------- *)  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
173  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
174  | 
type arguments =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
175  | 
  {
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
176  | 
(* just passed unchanged from 'params': *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
177  | 
maxvars : int,  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
178  | 
(* whether to use 'make_equality' or 'make_def_equality': *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
179  | 
def_eq : bool,  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
180  | 
(* the following may change during the translation: *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
181  | 
next_idx : int,  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
182  | 
bounds : interpretation list,  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
183  | 
wellformed: prop_formula  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
184  | 
};  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
185  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
186  | 
structure Data = Theory_Data  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
187  | 
(  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
188  | 
type T =  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
189  | 
    {interpreters: (string * (Proof.context -> model -> arguments -> term ->
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
190  | 
(interpretation * model * arguments) option)) list,  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
191  | 
printers: (string * (Proof.context -> model -> typ -> interpretation ->  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
192  | 
(int -> bool) -> term option)) list,  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
193  | 
parameters: string Symtab.table};  | 
| 46960 | 194  | 
  val empty = {interpreters = [], printers = [], parameters = Symtab.empty};
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
195  | 
val extend = I;  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
196  | 
fun merge  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
197  | 
    ({interpreters = in1, printers = pr1, parameters = pa1},
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
198  | 
     {interpreters = in2, printers = pr2, parameters = pa2}) : T =
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
199  | 
    {interpreters = AList.merge (op =) (K true) (in1, in2),
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
200  | 
printers = AList.merge (op =) (K true) (pr1, pr2),  | 
| 
41472
 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 
wenzelm 
parents: 
41471 
diff
changeset
 | 
201  | 
parameters = Symtab.merge (op =) (pa1, pa2)};  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
202  | 
);  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
203  | 
|
| 42361 | 204  | 
val get_data = Data.get o Proof_Context.theory_of;  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
205  | 
|
| 14350 | 206  | 
|
207  | 
(* ------------------------------------------------------------------------- *)  | 
|
| 
15334
 
d5a92997dc1b
exception CANNOT_INTERPRET removed (not needed anymore since the stlc_interpreter can interpret any term)
 
webertj 
parents: 
15333 
diff
changeset
 | 
208  | 
(* interpret: interprets the term 't' using a suitable interpreter; returns *)  | 
| 
 
d5a92997dc1b
exception CANNOT_INTERPRET removed (not needed anymore since the stlc_interpreter can interpret any term)
 
webertj 
parents: 
15333 
diff
changeset
 | 
209  | 
(* the interpretation and a (possibly extended) model that keeps *)  | 
| 
 
d5a92997dc1b
exception CANNOT_INTERPRET removed (not needed anymore since the stlc_interpreter can interpret any term)
 
webertj 
parents: 
15333 
diff
changeset
 | 
210  | 
(* track of the interpretation of subterms *)  | 
| 14350 | 211  | 
(* ------------------------------------------------------------------------- *)  | 
212  | 
||
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
213  | 
fun interpret ctxt model args t =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
214  | 
case get_first (fn (_, f) => f ctxt model args t)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
215  | 
(#interpreters (get_data ctxt)) of  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
216  | 
    NONE => raise REFUTE ("interpret",
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
217  | 
"no interpreter for term " ^ quote (Syntax.string_of_term ctxt t))  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
218  | 
| SOME x => x;  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
219  | 
|
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
220  | 
(* ------------------------------------------------------------------------- *)  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
221  | 
(* print: converts the interpretation 'intr', which must denote a term of *)  | 
| 
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
222  | 
(* type 'T', into a term using a suitable printer *)  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
223  | 
(* ------------------------------------------------------------------------- *)  | 
| 14350 | 224  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
225  | 
fun print ctxt model T intr assignment =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
226  | 
case get_first (fn (_, f) => f ctxt model T intr assignment)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
227  | 
(#printers (get_data ctxt)) of  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
228  | 
    NONE => raise REFUTE ("print",
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
229  | 
"no printer for type " ^ quote (Syntax.string_of_typ ctxt T))  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
230  | 
| SOME x => x;  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
231  | 
|
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
232  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
233  | 
(* print_model: turns the model into a string, using a fixed interpretation *)  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
234  | 
(* (given by an assignment for Boolean variables) and suitable *)  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
235  | 
(* printers *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
236  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
237  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
238  | 
fun print_model ctxt model assignment =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
239  | 
let  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
240  | 
val (typs, terms) = model  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
241  | 
val typs_msg =  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
242  | 
if null typs then  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
243  | 
"empty universe (no type variables in term)\n"  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
244  | 
else  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
245  | 
"Size of types: " ^ commas (map (fn (T, i) =>  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
246  | 
Syntax.string_of_typ ctxt T ^ ": " ^ string_of_int i) typs) ^ "\n"  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
247  | 
val show_consts_msg =  | 
| 
39050
 
600de0485859
turned show_consts into proper configuration option;
 
wenzelm 
parents: 
39049 
diff
changeset
 | 
248  | 
if not (Config.get ctxt show_consts) andalso Library.exists (is_Const o fst) terms then  | 
| 
 
600de0485859
turned show_consts into proper configuration option;
 
wenzelm 
parents: 
39049 
diff
changeset
 | 
249  | 
"enable \"show_consts\" to show the interpretation of constants\n"  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
250  | 
else  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
251  | 
""  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
252  | 
val terms_msg =  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
253  | 
if null terms then  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
254  | 
"empty interpretation (no free variables in term)\n"  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
255  | 
else  | 
| 32952 | 256  | 
cat_lines (map_filter (fn (t, intr) =>  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
257  | 
(* print constants only if 'show_consts' is true *)  | 
| 
39050
 
600de0485859
turned show_consts into proper configuration option;
 
wenzelm 
parents: 
39049 
diff
changeset
 | 
258  | 
if Config.get ctxt show_consts orelse not (is_Const t) then  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
259  | 
SOME (Syntax.string_of_term ctxt t ^ ": " ^  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
260  | 
Syntax.string_of_term ctxt  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
261  | 
(print ctxt model (Term.type_of t) intr assignment))  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
262  | 
else  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
263  | 
NONE) terms) ^ "\n"  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
264  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
265  | 
typs_msg ^ show_consts_msg ^ terms_msg  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
266  | 
end;  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
267  | 
|
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
268  | 
|
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
269  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
270  | 
(* PARAMETER MANAGEMENT *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
271  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
272  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
273  | 
fun add_interpreter name f = Data.map (fn {interpreters, printers, parameters} =>
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
274  | 
case AList.lookup (op =) interpreters name of  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
275  | 
    NONE => {interpreters = (name, f) :: interpreters,
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
276  | 
printers = printers, parameters = parameters}  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
277  | 
  | SOME _ => error ("Interpreter " ^ name ^ " already declared"));
 | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
278  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
279  | 
fun add_printer name f = Data.map (fn {interpreters, printers, parameters} =>
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
280  | 
case AList.lookup (op =) printers name of  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
281  | 
    NONE => {interpreters = interpreters,
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
282  | 
printers = (name, f) :: printers, parameters = parameters}  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
283  | 
  | SOME _ => error ("Printer " ^ name ^ " already declared"));
 | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
284  | 
|
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
285  | 
(* ------------------------------------------------------------------------- *)  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
286  | 
(* set_default_param: stores the '(name, value)' pair in Data's *)  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
287  | 
(* parameter table *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
288  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
289  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
290  | 
fun set_default_param (name, value) = Data.map  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
291  | 
  (fn {interpreters, printers, parameters} =>
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
292  | 
    {interpreters = interpreters, printers = printers,
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
293  | 
parameters = Symtab.update (name, value) parameters});  | 
| 14350 | 294  | 
|
295  | 
(* ------------------------------------------------------------------------- *)  | 
|
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
296  | 
(* get_default_param: retrieves the value associated with 'name' from *)  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
297  | 
(* Data's parameter table *)  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
298  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
299  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
300  | 
val get_default_param = Symtab.lookup o #parameters o get_data;  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
301  | 
|
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
302  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
303  | 
(* get_default_params: returns a list of all '(name, value)' pairs that are *)  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
304  | 
(* stored in Data's parameter table *)  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
305  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
306  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
307  | 
val get_default_params = Symtab.dest o #parameters o get_data;  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
308  | 
|
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
309  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
310  | 
(* actual_params: takes a (possibly empty) list 'params' of parameters that *)  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
311  | 
(* override the default parameters currently specified, and *)  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
312  | 
(* returns a record that can be passed to 'find_model'. *)  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
313  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
314  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
315  | 
fun actual_params ctxt override =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
316  | 
let  | 
| 
34120
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
317  | 
fun read_bool (parms, name) =  | 
| 
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
318  | 
case AList.lookup (op =) parms name of  | 
| 
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
319  | 
SOME "true" => true  | 
| 
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
320  | 
| SOME "false" => false  | 
| 
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
321  | 
      | SOME s => error ("parameter " ^ quote name ^
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
322  | 
" (value is " ^ quote s ^ ") must be \"true\" or \"false\"")  | 
| 
34120
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
323  | 
      | NONE   => error ("parameter " ^ quote name ^
 | 
| 
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
324  | 
" must be assigned a value")  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
325  | 
fun read_int (parms, name) =  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
326  | 
case AList.lookup (op =) parms name of  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
327  | 
SOME s =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
328  | 
(case Int.fromString s of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
329  | 
SOME i => i  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
330  | 
          | NONE   => error ("parameter " ^ quote name ^
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
331  | 
" (value is " ^ quote s ^ ") must be an integer value"))  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
332  | 
      | NONE => error ("parameter " ^ quote name ^
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
333  | 
" must be assigned a value")  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
334  | 
fun read_string (parms, name) =  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
335  | 
case AList.lookup (op =) parms name of  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
336  | 
SOME s => s  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
337  | 
      | NONE => error ("parameter " ^ quote name ^
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
338  | 
" must be assigned a value")  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
339  | 
(* 'override' first, defaults last: *)  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
340  | 
val allparams = override @ get_default_params ctxt  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
341  | 
val minsize = read_int (allparams, "minsize")  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
342  | 
val maxsize = read_int (allparams, "maxsize")  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
343  | 
val maxvars = read_int (allparams, "maxvars")  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
344  | 
val maxtime = read_int (allparams, "maxtime")  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
345  | 
val satsolver = read_string (allparams, "satsolver")  | 
| 
34120
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
346  | 
val no_assms = read_bool (allparams, "no_assms")  | 
| 
30314
 
853778f6ef7d
Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
 
blanchet 
parents: 
30312 
diff
changeset
 | 
347  | 
val expect = the_default "" (AList.lookup (op =) allparams "expect")  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
348  | 
(* all remaining parameters of the form "string=int" are collected in *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
349  | 
(* 'sizes' *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
350  | 
(* TODO: it is currently not possible to specify a size for a type *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
351  | 
(* whose name is one of the other parameters (e.g. 'maxvars') *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
352  | 
(* (string * int) list *)  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
353  | 
val sizes = map_filter  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
354  | 
(fn (name, value) => Option.map (pair name) (Int.fromString value))  | 
| 33317 | 355  | 
(filter (fn (name, _) => name<>"minsize" andalso name<>"maxsize"  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
356  | 
andalso name<>"maxvars" andalso name<>"maxtime"  | 
| 
34120
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
357  | 
andalso name<>"satsolver" andalso name<>"no_assms") allparams)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
358  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
359  | 
    {sizes=sizes, minsize=minsize, maxsize=maxsize, maxvars=maxvars,
 | 
| 
34120
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
360  | 
maxtime=maxtime, satsolver=satsolver, no_assms=no_assms, expect=expect}  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
361  | 
end;  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
362  | 
|
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
363  | 
|
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
364  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
365  | 
(* TRANSLATION HOL -> PROPOSITIONAL LOGIC, BOOLEAN ASSIGNMENT -> MODEL *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
366  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
367  | 
|
| 58120 | 368  | 
fun typ_of_dtyp _ typ_assoc (Old_Datatype_Aux.DtTFree a) =  | 
369  | 
the (AList.lookup (op =) typ_assoc (Old_Datatype_Aux.DtTFree a))  | 
|
370  | 
| typ_of_dtyp descr typ_assoc (Old_Datatype_Aux.DtType (s, Us)) =  | 
|
| 
55891
 
d1a9b07783ab
support 'datatype_new'-defined datatypes in Nitpick + better support for 'codatatype's
 
blanchet 
parents: 
55507 
diff
changeset
 | 
371  | 
Type (s, map (typ_of_dtyp descr typ_assoc) Us)  | 
| 58120 | 372  | 
| typ_of_dtyp descr typ_assoc (Old_Datatype_Aux.DtRec i) =  | 
| 
55891
 
d1a9b07783ab
support 'datatype_new'-defined datatypes in Nitpick + better support for 'codatatype's
 
blanchet 
parents: 
55507 
diff
changeset
 | 
373  | 
let val (s, ds, _) = the (AList.lookup (op =) descr i) in  | 
| 
 
d1a9b07783ab
support 'datatype_new'-defined datatypes in Nitpick + better support for 'codatatype's
 
blanchet 
parents: 
55507 
diff
changeset
 | 
374  | 
Type (s, map (typ_of_dtyp descr typ_assoc) ds)  | 
| 
 
d1a9b07783ab
support 'datatype_new'-defined datatypes in Nitpick + better support for 'codatatype's
 
blanchet 
parents: 
55507 
diff
changeset
 | 
375  | 
end  | 
| 
 
d1a9b07783ab
support 'datatype_new'-defined datatypes in Nitpick + better support for 'codatatype's
 
blanchet 
parents: 
55507 
diff
changeset
 | 
376  | 
|
| 
54757
 
4960647932ec
use 'prop' rather than 'bool' systematically in Isar reconstruction code
 
blanchet 
parents: 
54556 
diff
changeset
 | 
377  | 
val close_form = ATP_Util.close_form  | 
| 
43085
 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 
blanchet 
parents: 
42680 
diff
changeset
 | 
378  | 
val specialize_type = ATP_Util.specialize_type  | 
| 
21985
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
379  | 
|
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
380  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
381  | 
(* is_const_of_class: returns 'true' iff 'Const (s, T)' is a constant that *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
382  | 
(* denotes membership to an axiomatic type class *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
383  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
384  | 
|
| 46096 | 385  | 
fun is_const_of_class thy (s, _) =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
386  | 
let  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
387  | 
val class_const_names = map Logic.const_of_class (Sign.all_classes thy)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
388  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
389  | 
(* I'm not quite sure if checking the name 's' is sufficient, *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
390  | 
(* or if we should also check the type 'T'. *)  | 
| 
36692
 
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
 
haftmann 
parents: 
36555 
diff
changeset
 | 
391  | 
member (op =) class_const_names s  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
392  | 
end;  | 
| 
21985
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
393  | 
|
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
394  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
395  | 
(* is_IDT_constructor: returns 'true' iff 'Const (s, T)' is the constructor *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
396  | 
(* of an inductive datatype in 'thy' *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
397  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
398  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
399  | 
fun is_IDT_constructor thy (s, T) =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
400  | 
(case body_type T of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
401  | 
Type (s', _) =>  | 
| 
58129
 
3ec65a7f2b50
ported Refute to use new datatypes when possible
 
blanchet 
parents: 
58120 
diff
changeset
 | 
402  | 
(case BNF_LFP_Compat.get_constrs thy s' of  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
403  | 
SOME constrs =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
404  | 
List.exists (fn (cname, cty) =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
405  | 
cname = s andalso Sign.typ_instance thy (T, cty)) constrs  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
406  | 
| NONE => false)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
407  | 
| _ => false);  | 
| 
21985
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
408  | 
|
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
409  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
410  | 
(* is_IDT_recursor: returns 'true' iff 'Const (s, T)' is the recursion *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
411  | 
(* operator of an inductive datatype in 'thy' *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
412  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
413  | 
|
| 46096 | 414  | 
fun is_IDT_recursor thy (s, _) =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
415  | 
let  | 
| 
58354
 
04ac60da613e
support (finite values of) codatatypes in Quickcheck
 
blanchet 
parents: 
58322 
diff
changeset
 | 
416  | 
val rec_names = Symtab.fold (append o #rec_names o snd) (BNF_LFP_Compat.get_all thy []) []  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
417  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
418  | 
(* I'm not quite sure if checking the name 's' is sufficient, *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
419  | 
(* or if we should also check the type 'T'. *)  | 
| 
36692
 
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
 
haftmann 
parents: 
36555 
diff
changeset
 | 
420  | 
member (op =) rec_names s  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
421  | 
end;  | 
| 
21985
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
422  | 
|
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
423  | 
(* ------------------------------------------------------------------------- *)  | 
| 
30275
 
381ce8d88cb8
Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
 
blanchet 
parents: 
30242 
diff
changeset
 | 
424  | 
(* norm_rhs: maps f ?t1 ... ?tn == rhs to %t1...tn. rhs *)  | 
| 
 
381ce8d88cb8
Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
 
blanchet 
parents: 
30242 
diff
changeset
 | 
425  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
381ce8d88cb8
Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
 
blanchet 
parents: 
30242 
diff
changeset
 | 
426  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
427  | 
fun norm_rhs eqn =  | 
| 
30275
 
381ce8d88cb8
Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
 
blanchet 
parents: 
30242 
diff
changeset
 | 
428  | 
let  | 
| 
 
381ce8d88cb8
Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
 
blanchet 
parents: 
30242 
diff
changeset
 | 
429  | 
fun lambda (v as Var ((x, _), T)) t = Abs (x, T, abstract_over (v, t))  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
430  | 
      | lambda v t = raise TERM ("lambda", [v, t])
 | 
| 
30275
 
381ce8d88cb8
Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
 
blanchet 
parents: 
30242 
diff
changeset
 | 
431  | 
val (lhs, rhs) = Logic.dest_equals eqn  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
432  | 
val (_, args) = Term.strip_comb lhs  | 
| 
30275
 
381ce8d88cb8
Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
 
blanchet 
parents: 
30242 
diff
changeset
 | 
433  | 
in  | 
| 
 
381ce8d88cb8
Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
 
blanchet 
parents: 
30242 
diff
changeset
 | 
434  | 
fold lambda (rev args) rhs  | 
| 
 
381ce8d88cb8
Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
 
blanchet 
parents: 
30242 
diff
changeset
 | 
435  | 
end  | 
| 
 
381ce8d88cb8
Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
 
blanchet 
parents: 
30242 
diff
changeset
 | 
436  | 
|
| 
 
381ce8d88cb8
Reintroduced previous changes: Made "Refute.norm_rhs" public and simplified the configuration of the BerkMin and zChaff SAT solvers.
 
blanchet 
parents: 
30242 
diff
changeset
 | 
437  | 
(* ------------------------------------------------------------------------- *)  | 
| 37405 | 438  | 
(* get_def: looks up the definition of a constant *)  | 
| 
21985
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
439  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
440  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
441  | 
fun get_def thy (s, T) =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
442  | 
let  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
443  | 
fun get_def_ax [] = NONE  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
444  | 
| get_def_ax ((axname, ax) :: axioms) =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
445  | 
(let  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
446  | 
val (lhs, _) = Logic.dest_equals ax (* equations only *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
447  | 
val c = Term.head_of lhs  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
448  | 
val (s', T') = Term.dest_Const c  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
449  | 
in  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
450  | 
if s=s' then  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
451  | 
let  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
452  | 
val typeSubs = Sign.typ_match thy (T', T) Vartab.empty  | 
| 61770 | 453  | 
val ax' = Envir.subst_term_types typeSubs ax  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
454  | 
val rhs = norm_rhs ax'  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
455  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
456  | 
SOME (axname, rhs)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
457  | 
end  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
458  | 
else  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
459  | 
get_def_ax axioms  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
460  | 
end handle ERROR _ => get_def_ax axioms  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
461  | 
| TERM _ => get_def_ax axioms  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
462  | 
| Type.TYPE_MATCH => get_def_ax axioms)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
463  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
464  | 
get_def_ax (Theory.all_axioms_of thy)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
465  | 
end;  | 
| 
21985
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
466  | 
|
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
467  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
468  | 
(* get_typedef: looks up the definition of a type, as created by "typedef" *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
469  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
470  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
471  | 
fun get_typedef thy T =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
472  | 
let  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
473  | 
fun get_typedef_ax [] = NONE  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
474  | 
| get_typedef_ax ((axname, ax) :: axioms) =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
475  | 
(let  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
476  | 
fun type_of_type_definition (Const (s', T')) =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
477  | 
                  if s'= @{const_name type_definition} then
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
478  | 
SOME T'  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
479  | 
else  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
480  | 
NONE  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
481  | 
| type_of_type_definition (Free _) = NONE  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
482  | 
| type_of_type_definition (Var _) = NONE  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
483  | 
| type_of_type_definition (Bound _) = NONE  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
484  | 
| type_of_type_definition (Abs (_, _, body)) =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
485  | 
type_of_type_definition body  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
486  | 
| type_of_type_definition (t1 $ t2) =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
487  | 
(case type_of_type_definition t1 of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
488  | 
SOME x => SOME x  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
489  | 
| NONE => type_of_type_definition t2)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
490  | 
in  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
491  | 
case type_of_type_definition ax of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
492  | 
SOME T' =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
493  | 
let  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
494  | 
val T'' = domain_type (domain_type T')  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
495  | 
val typeSubs = Sign.typ_match thy (T'', T) Vartab.empty  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
496  | 
in  | 
| 61770 | 497  | 
SOME (axname, Envir.subst_term_types typeSubs ax)  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
498  | 
end  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
499  | 
| NONE => get_typedef_ax axioms  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
500  | 
end handle ERROR _ => get_typedef_ax axioms  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
501  | 
| TERM _ => get_typedef_ax axioms  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
502  | 
| Type.TYPE_MATCH => get_typedef_ax axioms)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
503  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
504  | 
get_typedef_ax (Theory.all_axioms_of thy)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
505  | 
end;  | 
| 
21985
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
506  | 
|
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
507  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
508  | 
(* get_classdef: looks up the defining axiom for an axiomatic type class, as *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
509  | 
(* created by the "axclass" command *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
510  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
511  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
512  | 
fun get_classdef thy class =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
513  | 
let  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
514  | 
val axname = class ^ "_class_def"  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
515  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
516  | 
Option.map (pair axname)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
517  | 
(AList.lookup (op =) (Theory.all_axioms_of thy) axname)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
518  | 
end;  | 
| 
21985
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
519  | 
|
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
520  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
521  | 
(* unfold_defs: unfolds all defined constants in a term 't', beta-eta *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
522  | 
(* normalizes the result term; certain constants are not *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
523  | 
(* unfolded (cf. 'collect_axioms' and the various interpreters *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
524  | 
(* below): if the interpretation respects a definition anyway, *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
525  | 
(* that definition does not need to be unfolded *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
526  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
527  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
528  | 
(* Note: we could intertwine unfolding of constants and beta-(eta-) *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
529  | 
(* normalization; this would save some unfolding for terms where *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
530  | 
(* constants are eliminated by beta-reduction (e.g. 'K c1 c2'). On *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
531  | 
(* the other hand, this would cause additional work for terms where *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
532  | 
(* constants are duplicated by beta-reduction (e.g. 'S c1 c2 c3'). *)  | 
| 
21985
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
533  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
534  | 
fun unfold_defs thy t =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
535  | 
let  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
536  | 
fun unfold_loop t =  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
537  | 
case t of  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
538  | 
(* Pure *)  | 
| 56245 | 539  | 
        Const (@{const_name Pure.all}, _) => t
 | 
540  | 
      | Const (@{const_name Pure.eq}, _) => t
 | 
|
541  | 
      | Const (@{const_name Pure.imp}, _) => t
 | 
|
| 56243 | 542  | 
      | Const (@{const_name Pure.type}, _) => t  (* axiomatic type classes *)
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
543  | 
(* HOL *)  | 
| 
29802
 
d201a838d0f7
Make some Refute functions public so I can use them in Nitpick,
 
blanchet 
parents: 
29288 
diff
changeset
 | 
544  | 
      | Const (@{const_name Trueprop}, _) => t
 | 
| 
 
d201a838d0f7
Make some Refute functions public so I can use them in Nitpick,
 
blanchet 
parents: 
29288 
diff
changeset
 | 
545  | 
      | Const (@{const_name Not}, _) => t
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
546  | 
| (* redundant, since 'True' is also an IDT constructor *)  | 
| 
29802
 
d201a838d0f7
Make some Refute functions public so I can use them in Nitpick,
 
blanchet 
parents: 
29288 
diff
changeset
 | 
547  | 
        Const (@{const_name True}, _) => t
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
548  | 
| (* redundant, since 'False' is also an IDT constructor *)  | 
| 
29802
 
d201a838d0f7
Make some Refute functions public so I can use them in Nitpick,
 
blanchet 
parents: 
29288 
diff
changeset
 | 
549  | 
        Const (@{const_name False}, _) => t
 | 
| 
 
d201a838d0f7
Make some Refute functions public so I can use them in Nitpick,
 
blanchet 
parents: 
29288 
diff
changeset
 | 
550  | 
      | Const (@{const_name undefined}, _) => t
 | 
| 
 
d201a838d0f7
Make some Refute functions public so I can use them in Nitpick,
 
blanchet 
parents: 
29288 
diff
changeset
 | 
551  | 
      | Const (@{const_name The}, _) => t
 | 
| 
29820
 
07f53494cf20
Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
 
blanchet 
parents: 
29802 
diff
changeset
 | 
552  | 
      | Const (@{const_name Hilbert_Choice.Eps}, _) => t
 | 
| 
29802
 
d201a838d0f7
Make some Refute functions public so I can use them in Nitpick,
 
blanchet 
parents: 
29288 
diff
changeset
 | 
553  | 
      | Const (@{const_name All}, _) => t
 | 
| 
 
d201a838d0f7
Make some Refute functions public so I can use them in Nitpick,
 
blanchet 
parents: 
29288 
diff
changeset
 | 
554  | 
      | Const (@{const_name Ex}, _) => t
 | 
| 
38864
 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 
haftmann 
parents: 
38795 
diff
changeset
 | 
555  | 
      | Const (@{const_name HOL.eq}, _) => t
 | 
| 
38795
 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 
haftmann 
parents: 
38786 
diff
changeset
 | 
556  | 
      | Const (@{const_name HOL.conj}, _) => t
 | 
| 
 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 
haftmann 
parents: 
38786 
diff
changeset
 | 
557  | 
      | Const (@{const_name HOL.disj}, _) => t
 | 
| 
38786
 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 
haftmann 
parents: 
38553 
diff
changeset
 | 
558  | 
      | Const (@{const_name HOL.implies}, _) => t
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
559  | 
(* sets *)  | 
| 
29802
 
d201a838d0f7
Make some Refute functions public so I can use them in Nitpick,
 
blanchet 
parents: 
29288 
diff
changeset
 | 
560  | 
      | Const (@{const_name Collect}, _) => t
 | 
| 37677 | 561  | 
      | Const (@{const_name Set.member}, _) => t
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
562  | 
(* other optimizations *)  | 
| 
29820
 
07f53494cf20
Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
 
blanchet 
parents: 
29802 
diff
changeset
 | 
563  | 
      | Const (@{const_name Finite_Set.card}, _) => t
 | 
| 
 
07f53494cf20
Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
 
blanchet 
parents: 
29802 
diff
changeset
 | 
564  | 
      | Const (@{const_name Finite_Set.finite}, _) => t
 | 
| 37388 | 565  | 
      | Const (@{const_name Orderings.less}, Type ("fun", [@{typ nat},
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
566  | 
          Type ("fun", [@{typ nat}, @{typ bool}])])) => t
 | 
| 37388 | 567  | 
      | Const (@{const_name Groups.plus}, Type ("fun", [@{typ nat},
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
568  | 
          Type ("fun", [@{typ nat}, @{typ nat}])])) => t
 | 
| 37388 | 569  | 
      | Const (@{const_name Groups.minus}, Type ("fun", [@{typ nat},
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
570  | 
          Type ("fun", [@{typ nat}, @{typ nat}])])) => t
 | 
| 37388 | 571  | 
      | Const (@{const_name Groups.times}, Type ("fun", [@{typ nat},
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
572  | 
          Type ("fun", [@{typ nat}, @{typ nat}])])) => t
 | 
| 56252 | 573  | 
      | Const (@{const_name append}, _) => t
 | 
| 
29820
 
07f53494cf20
Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
 
blanchet 
parents: 
29802 
diff
changeset
 | 
574  | 
      | Const (@{const_name fst}, _) => t
 | 
| 
 
07f53494cf20
Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
 
blanchet 
parents: 
29802 
diff
changeset
 | 
575  | 
      | Const (@{const_name snd}, _) => t
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
576  | 
(* simply-typed lambda calculus *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
577  | 
| Const (s, T) =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
578  | 
(if is_IDT_constructor thy (s, T)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
579  | 
orelse is_IDT_recursor thy (s, T) then  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
580  | 
t (* do not unfold IDT constructors/recursors *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
581  | 
(* unfold the constant if there is a defining equation *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
582  | 
else  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
583  | 
case get_def thy (s, T) of  | 
| 46096 | 584  | 
SOME ((*axname*) _, rhs) =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
585  | 
(* Note: if the term to be unfolded (i.e. 'Const (s, T)') *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
586  | 
(* occurs on the right-hand side of the equation, i.e. in *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
587  | 
(* 'rhs', we must not use this equation to unfold, because *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
588  | 
(* that would loop. Here would be the right place to *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
589  | 
(* check this. However, getting this really right seems *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
590  | 
(* difficult because the user may state arbitrary axioms, *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
591  | 
(* which could interact with overloading to create loops. *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
592  | 
              ((*tracing (" unfolding: " ^ axname);*)
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
593  | 
unfold_loop rhs)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
594  | 
| NONE => t)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
595  | 
| Free _ => t  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
596  | 
| Var _ => t  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
597  | 
| Bound _ => t  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
598  | 
| Abs (s, T, body) => Abs (s, T, unfold_loop body)  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
599  | 
| t1 $ t2 => (unfold_loop t1) $ (unfold_loop t2)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
600  | 
val result = Envir.beta_eta_contract (unfold_loop t)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
601  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
602  | 
result  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
603  | 
end;  | 
| 
21985
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
604  | 
|
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
605  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
606  | 
(* collect_axioms: collects (monomorphic, universally quantified, unfolded *)  | 
| 
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
607  | 
(* versions of) all HOL axioms that are relevant w.r.t 't' *)  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
608  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
609  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
610  | 
(* Note: to make the collection of axioms more easily extensible, this *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
611  | 
(* function could be based on user-supplied "axiom collectors", *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
612  | 
(* similar to 'interpret'/interpreters or 'print'/printers *)  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
613  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
614  | 
(* Note: currently we use "inverse" functions to the definitional *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
615  | 
(* mechanisms provided by Isabelle/HOL, e.g. for "axclass", *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
616  | 
(* "typedef", "definition". A more general approach could consider *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
617  | 
(* *every* axiom of the theory and collect it if it has a constant/ *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
618  | 
(* type/typeclass in common with the term 't'. *)  | 
| 
21985
 
acfb13e8819e
constants are unfolded, universal quantifiers are stripped, some minor changes
 
webertj 
parents: 
21931 
diff
changeset
 | 
619  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
620  | 
(* Which axioms are "relevant" for a particular term/type goes hand in *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
621  | 
(* hand with the interpretation of that term/type by its interpreter (see *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
622  | 
(* way below): if the interpretation respects an axiom anyway, the axiom *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
623  | 
(* does not need to be added as a constraint here. *)  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
624  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
625  | 
(* To avoid collecting the same axiom multiple times, we use an *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
626  | 
(* accumulator 'axs' which contains all axioms collected so far. *)  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
627  | 
|
| 56256 | 628  | 
local  | 
629  | 
||
630  | 
fun get_axiom thy xname =  | 
|
631  | 
Name_Space.check (Context.Theory thy) (Theory.axiom_table thy) (xname, Position.none);  | 
|
632  | 
||
633  | 
val the_eq_trivial = get_axiom @{theory HOL} "the_eq_trivial";
 | 
|
634  | 
val someI = get_axiom @{theory Hilbert_Choice} "someI";
 | 
|
635  | 
||
636  | 
in  | 
|
637  | 
||
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
638  | 
fun collect_axioms ctxt t =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
639  | 
let  | 
| 42361 | 640  | 
val thy = Proof_Context.theory_of ctxt  | 
| 32950 | 641  | 
val _ = tracing "Adding axioms..."  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
642  | 
fun collect_this_axiom (axname, ax) axs =  | 
| 33246 | 643  | 
let  | 
644  | 
val ax' = unfold_defs thy ax  | 
|
645  | 
in  | 
|
646  | 
if member (op aconv) axs ax' then axs  | 
|
647  | 
else (tracing axname; collect_term_axioms ax' (ax' :: axs))  | 
|
648  | 
end  | 
|
649  | 
and collect_sort_axioms T axs =  | 
|
650  | 
let  | 
|
651  | 
val sort =  | 
|
652  | 
(case T of  | 
|
653  | 
TFree (_, sort) => sort  | 
|
654  | 
| TVar (_, sort) => sort  | 
|
655  | 
          | _ => raise REFUTE ("collect_axioms",
 | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
656  | 
"type " ^ Syntax.string_of_typ ctxt T ^ " is not a variable"))  | 
| 33246 | 657  | 
(* obtain axioms for all superclasses *)  | 
658  | 
val superclasses = sort @ maps (Sign.super_classes thy) sort  | 
|
659  | 
(* merely an optimization, because 'collect_this_axiom' disallows *)  | 
|
660  | 
(* duplicate axioms anyway: *)  | 
|
661  | 
val superclasses = distinct (op =) superclasses  | 
|
662  | 
val class_axioms = maps (fn class => map (fn ax =>  | 
|
663  | 
          ("<" ^ class ^ ">", Thm.prop_of ax))
 | 
|
| 
51685
 
385ef6706252
more standard module name Axclass (according to file name);
 
wenzelm 
parents: 
51557 
diff
changeset
 | 
664  | 
(#axioms (Axclass.get_info thy class) handle ERROR _ => []))  | 
| 33246 | 665  | 
superclasses  | 
666  | 
(* replace the (at most one) schematic type variable in each axiom *)  | 
|
667  | 
(* by the actual type 'T' *)  | 
|
668  | 
val monomorphic_class_axioms = map (fn (axname, ax) =>  | 
|
669  | 
(case Term.add_tvars ax [] of  | 
|
670  | 
[] => (axname, ax)  | 
|
| 61770 | 671  | 
| [(idx, S)] => (axname, Envir.subst_term_types (Vartab.make [(idx, (S, T))]) ax)  | 
| 33246 | 672  | 
| _ =>  | 
673  | 
            raise REFUTE ("collect_axioms", "class axiom " ^ axname ^ " (" ^
 | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
674  | 
Syntax.string_of_term ctxt ax ^  | 
| 33246 | 675  | 
") contains more than one type variable")))  | 
676  | 
class_axioms  | 
|
677  | 
in  | 
|
678  | 
fold collect_this_axiom monomorphic_class_axioms axs  | 
|
679  | 
end  | 
|
680  | 
and collect_type_axioms T axs =  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
681  | 
case T of  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
682  | 
(* simple types *)  | 
| 56252 | 683  | 
        Type (@{type_name prop}, []) => axs
 | 
684  | 
      | Type (@{type_name fun}, [T1, T2]) => collect_type_axioms T2 (collect_type_axioms T1 axs)
 | 
|
| 
46098
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
685  | 
      | Type (@{type_name set}, [T1]) => collect_type_axioms T1 axs
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
686  | 
(* axiomatic type classes *)  | 
| 56252 | 687  | 
      | Type (@{type_name itself}, [T1]) => collect_type_axioms T1 axs
 | 
| 33246 | 688  | 
| Type (s, Ts) =>  | 
| 
58354
 
04ac60da613e
support (finite values of) codatatypes in Quickcheck
 
blanchet 
parents: 
58322 
diff
changeset
 | 
689  | 
(case BNF_LFP_Compat.get_info thy [] s of  | 
| 46096 | 690  | 
SOME _ => (* inductive datatype *)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
691  | 
(* only collect relevant type axioms for the argument types *)  | 
| 33246 | 692  | 
fold collect_type_axioms Ts axs  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
693  | 
| NONE =>  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
694  | 
(case get_typedef thy T of  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
695  | 
SOME (axname, ax) =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
696  | 
collect_this_axiom (axname, ax) axs  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
697  | 
| NONE =>  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
698  | 
(* unspecified type, perhaps introduced with "typedecl" *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
699  | 
(* at least collect relevant type axioms for the argument types *)  | 
| 33246 | 700  | 
fold collect_type_axioms Ts axs))  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
701  | 
(* axiomatic type classes *)  | 
| 33246 | 702  | 
| TFree _ => collect_sort_axioms T axs  | 
703  | 
(* axiomatic type classes *)  | 
|
704  | 
| TVar _ => collect_sort_axioms T axs  | 
|
705  | 
and collect_term_axioms t axs =  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
706  | 
case t of  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
707  | 
(* Pure *)  | 
| 56245 | 708  | 
        Const (@{const_name Pure.all}, _) => axs
 | 
709  | 
      | Const (@{const_name Pure.eq}, _) => axs
 | 
|
710  | 
      | Const (@{const_name Pure.imp}, _) => axs
 | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
711  | 
(* axiomatic type classes *)  | 
| 56243 | 712  | 
      | Const (@{const_name Pure.type}, T) => collect_type_axioms T axs
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
713  | 
(* HOL *)  | 
| 
29802
 
d201a838d0f7
Make some Refute functions public so I can use them in Nitpick,
 
blanchet 
parents: 
29288 
diff
changeset
 | 
714  | 
      | Const (@{const_name Trueprop}, _) => axs
 | 
| 
 
d201a838d0f7
Make some Refute functions public so I can use them in Nitpick,
 
blanchet 
parents: 
29288 
diff
changeset
 | 
715  | 
      | Const (@{const_name Not}, _) => axs
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
716  | 
(* redundant, since 'True' is also an IDT constructor *)  | 
| 
29802
 
d201a838d0f7
Make some Refute functions public so I can use them in Nitpick,
 
blanchet 
parents: 
29288 
diff
changeset
 | 
717  | 
      | Const (@{const_name True}, _) => axs
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
718  | 
(* redundant, since 'False' is also an IDT constructor *)  | 
| 
29802
 
d201a838d0f7
Make some Refute functions public so I can use them in Nitpick,
 
blanchet 
parents: 
29288 
diff
changeset
 | 
719  | 
      | Const (@{const_name False}, _) => axs
 | 
| 33246 | 720  | 
      | Const (@{const_name undefined}, T) => collect_type_axioms T axs
 | 
| 
29802
 
d201a838d0f7
Make some Refute functions public so I can use them in Nitpick,
 
blanchet 
parents: 
29288 
diff
changeset
 | 
721  | 
      | Const (@{const_name The}, T) =>
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
722  | 
let  | 
| 56256 | 723  | 
            val ax = specialize_type thy (@{const_name The}, T) (#2 the_eq_trivial)
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
724  | 
in  | 
| 56256 | 725  | 
collect_this_axiom (#1 the_eq_trivial, ax) axs  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
726  | 
end  | 
| 
29820
 
07f53494cf20
Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
 
blanchet 
parents: 
29802 
diff
changeset
 | 
727  | 
      | Const (@{const_name Hilbert_Choice.Eps}, T) =>
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
728  | 
let  | 
| 56256 | 729  | 
            val ax = specialize_type thy (@{const_name Hilbert_Choice.Eps}, T) (#2 someI)
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
730  | 
in  | 
| 56256 | 731  | 
collect_this_axiom (#1 someI, ax) axs  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
732  | 
end  | 
| 33246 | 733  | 
      | Const (@{const_name All}, T) => collect_type_axioms T axs
 | 
734  | 
      | Const (@{const_name Ex}, T) => collect_type_axioms T axs
 | 
|
| 
38864
 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 
haftmann 
parents: 
38795 
diff
changeset
 | 
735  | 
      | Const (@{const_name HOL.eq}, T) => collect_type_axioms T axs
 | 
| 
38795
 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 
haftmann 
parents: 
38786 
diff
changeset
 | 
736  | 
      | Const (@{const_name HOL.conj}, _) => axs
 | 
| 
 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 
haftmann 
parents: 
38786 
diff
changeset
 | 
737  | 
      | Const (@{const_name HOL.disj}, _) => axs
 | 
| 
38786
 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 
haftmann 
parents: 
38553 
diff
changeset
 | 
738  | 
      | Const (@{const_name HOL.implies}, _) => axs
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
739  | 
(* sets *)  | 
| 33246 | 740  | 
      | Const (@{const_name Collect}, T) => collect_type_axioms T axs
 | 
| 37677 | 741  | 
      | Const (@{const_name Set.member}, T) => collect_type_axioms T axs
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
742  | 
(* other optimizations *)  | 
| 33246 | 743  | 
      | Const (@{const_name Finite_Set.card}, T) => collect_type_axioms T axs
 | 
| 
29820
 
07f53494cf20
Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
 
blanchet 
parents: 
29802 
diff
changeset
 | 
744  | 
      | Const (@{const_name Finite_Set.finite}, T) =>
 | 
| 33246 | 745  | 
collect_type_axioms T axs  | 
| 37388 | 746  | 
      | Const (@{const_name Orderings.less}, T as Type ("fun", [@{typ nat},
 | 
| 
38553
 
56965d8e4e11
use HOLogic.boolT and @{typ bool} more pervasively
 
haftmann 
parents: 
37677 
diff
changeset
 | 
747  | 
        Type ("fun", [@{typ nat}, @{typ bool}])])) =>
 | 
| 33246 | 748  | 
collect_type_axioms T axs  | 
| 37388 | 749  | 
      | Const (@{const_name Groups.plus}, T as Type ("fun", [@{typ nat},
 | 
750  | 
        Type ("fun", [@{typ nat}, @{typ nat}])])) =>
 | 
|
| 33246 | 751  | 
collect_type_axioms T axs  | 
| 37388 | 752  | 
      | Const (@{const_name Groups.minus}, T as Type ("fun", [@{typ nat},
 | 
753  | 
        Type ("fun", [@{typ nat}, @{typ nat}])])) =>
 | 
|
| 33246 | 754  | 
collect_type_axioms T axs  | 
| 37388 | 755  | 
      | Const (@{const_name Groups.times}, T as Type ("fun", [@{typ nat},
 | 
756  | 
        Type ("fun", [@{typ nat}, @{typ nat}])])) =>
 | 
|
| 33246 | 757  | 
collect_type_axioms T axs  | 
| 56252 | 758  | 
      | Const (@{const_name append}, T) => collect_type_axioms T axs
 | 
| 33246 | 759  | 
      | Const (@{const_name fst}, T) => collect_type_axioms T axs
 | 
760  | 
      | Const (@{const_name snd}, T) => collect_type_axioms T axs
 | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
761  | 
(* simply-typed lambda calculus *)  | 
| 
29802
 
d201a838d0f7
Make some Refute functions public so I can use them in Nitpick,
 
blanchet 
parents: 
29288 
diff
changeset
 | 
762  | 
| Const (s, T) =>  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
763  | 
if is_const_of_class thy (s, T) then  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
764  | 
(* axiomatic type classes: add "OFCLASS(?'a::c, c_class)" *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
765  | 
(* and the class definition *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
766  | 
let  | 
| 33246 | 767  | 
val class = Logic.class_of_const s  | 
| 
31943
 
5e960a0780a2
renamed inclass/Inclass to of_class/OfClass, in accordance to of_sort;
 
wenzelm 
parents: 
31784 
diff
changeset
 | 
768  | 
              val of_class = Logic.mk_of_class (TVar (("'a", 0), [class]), class)
 | 
| 33246 | 769  | 
val ax_in = SOME (specialize_type thy (s, T) of_class)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
770  | 
(* type match may fail due to sort constraints *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
771  | 
handle Type.TYPE_MATCH => NONE  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
772  | 
val ax_1 = Option.map (fn ax => (Syntax.string_of_term ctxt ax, ax)) ax_in  | 
| 33246 | 773  | 
val ax_2 = Option.map (apsnd (specialize_type thy (s, T))) (get_classdef thy class)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
774  | 
in  | 
| 33246 | 775  | 
collect_type_axioms T (fold collect_this_axiom (map_filter I [ax_1, ax_2]) axs)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
776  | 
end  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
777  | 
else if is_IDT_constructor thy (s, T)  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
778  | 
orelse is_IDT_recursor thy (s, T)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
779  | 
then  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
780  | 
(* only collect relevant type axioms *)  | 
| 33246 | 781  | 
collect_type_axioms T axs  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
782  | 
else  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
783  | 
(* other constants should have been unfolded, with some *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
784  | 
(* exceptions: e.g. Abs_xxx/Rep_xxx functions for *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
785  | 
(* typedefs, or type-class related constants *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
786  | 
(* only collect relevant type axioms *)  | 
| 33246 | 787  | 
collect_type_axioms T axs  | 
788  | 
| Free (_, T) => collect_type_axioms T axs  | 
|
789  | 
| Var (_, T) => collect_type_axioms T axs  | 
|
790  | 
| Bound _ => axs  | 
|
791  | 
| Abs (_, T, body) => collect_term_axioms body (collect_type_axioms T axs)  | 
|
792  | 
| t1 $ t2 => collect_term_axioms t2 (collect_term_axioms t1 axs)  | 
|
793  | 
val result = map close_form (collect_term_axioms t [])  | 
|
| 32950 | 794  | 
val _ = tracing " ...done."  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
795  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
796  | 
result  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
797  | 
end;  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
798  | 
|
| 56256 | 799  | 
end;  | 
800  | 
||
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
801  | 
(* ------------------------------------------------------------------------- *)  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
802  | 
(* ground_types: collects all ground types in a term (including argument *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
803  | 
(* types of other types), suppressing duplicates. Does not *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
804  | 
(* return function types, set types, non-recursive IDTs, or *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
805  | 
(* 'propT'. For IDTs, also the argument types of constructors *)  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
806  | 
(* and all mutually recursive IDTs are considered. *)  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
807  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
808  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
809  | 
fun ground_types ctxt t =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
810  | 
let  | 
| 42361 | 811  | 
val thy = Proof_Context.theory_of ctxt  | 
| 29272 | 812  | 
fun collect_types T acc =  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
813  | 
(case T of  | 
| 56252 | 814  | 
        Type (@{type_name fun}, [T1, T2]) => collect_types T1 (collect_types T2 acc)
 | 
815  | 
      | Type (@{type_name prop}, []) => acc
 | 
|
| 
46098
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
816  | 
      | Type (@{type_name set}, [T1]) => collect_types T1 acc
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
817  | 
| Type (s, Ts) =>  | 
| 
58354
 
04ac60da613e
support (finite values of) codatatypes in Quickcheck
 
blanchet 
parents: 
58322 
diff
changeset
 | 
818  | 
(case BNF_LFP_Compat.get_info thy [] s of  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
819  | 
SOME info => (* inductive datatype *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
820  | 
let  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
821  | 
val index = #index info  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
822  | 
val descr = #descr info  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
823  | 
val (_, typs, _) = the (AList.lookup (op =) descr index)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
824  | 
val typ_assoc = typs ~~ Ts  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
825  | 
(* sanity check: every element in 'dtyps' must be a *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
826  | 
(* 'DtTFree' *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
827  | 
val _ = if Library.exists (fn d =>  | 
| 58120 | 828  | 
case d of Old_Datatype_Aux.DtTFree _ => false | _ => true) typs then  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
829  | 
                  raise REFUTE ("ground_types", "datatype argument (for type "
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
830  | 
^ Syntax.string_of_typ ctxt T ^ ") is not a variable")  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
831  | 
else ()  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
832  | 
(* required for mutually recursive datatypes; those need to *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
833  | 
(* be added even if they are an instance of an otherwise non- *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
834  | 
(* recursive datatype *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
835  | 
fun collect_dtyp d acc =  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
836  | 
let  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
837  | 
val dT = typ_of_dtyp descr typ_assoc d  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
838  | 
in  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
839  | 
case d of  | 
| 58120 | 840  | 
Old_Datatype_Aux.DtTFree _ =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
841  | 
collect_types dT acc  | 
| 58120 | 842  | 
| Old_Datatype_Aux.DtType (_, ds) =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
843  | 
collect_types dT (fold_rev collect_dtyp ds acc)  | 
| 58120 | 844  | 
| Old_Datatype_Aux.DtRec i =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
845  | 
if member (op =) acc dT then  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
846  | 
acc (* prevent infinite recursion *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
847  | 
else  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
848  | 
let  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
849  | 
val (_, dtyps, dconstrs) = the (AList.lookup (op =) descr i)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
850  | 
(* if the current type is a recursive IDT (i.e. a depth *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
851  | 
(* is required), add it to 'acc' *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
852  | 
val acc_dT = if Library.exists (fn (_, ds) =>  | 
| 58120 | 853  | 
Library.exists Old_Datatype_Aux.is_rec_type ds) dconstrs then  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
854  | 
insert (op =) dT acc  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
855  | 
else acc  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
856  | 
(* collect argument types *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
857  | 
val acc_dtyps = fold_rev collect_dtyp dtyps acc_dT  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
858  | 
(* collect constructor types *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
859  | 
val acc_dconstrs = fold_rev collect_dtyp (maps snd dconstrs) acc_dtyps  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
860  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
861  | 
acc_dconstrs  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
862  | 
end  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
863  | 
end  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
864  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
865  | 
(* argument types 'Ts' could be added here, but they are also *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
866  | 
(* added by 'collect_dtyp' automatically *)  | 
| 58120 | 867  | 
collect_dtyp (Old_Datatype_Aux.DtRec index) acc  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
868  | 
end  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
869  | 
| NONE =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
870  | 
(* not an inductive datatype, e.g. defined via "typedef" or *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
871  | 
(* "typedecl" *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
872  | 
insert (op =) T (fold collect_types Ts acc))  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
873  | 
| TFree _ => insert (op =) T acc  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
874  | 
| TVar _ => insert (op =) T acc)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
875  | 
in  | 
| 29272 | 876  | 
fold_types collect_types t []  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
877  | 
end;  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
878  | 
|
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
879  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
880  | 
(* string_of_typ: (rather naive) conversion from types to strings, used to *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
881  | 
(* look up the size of a type in 'sizes'. Parameterized *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
882  | 
(* types with different parameters (e.g. "'a list" vs. "bool *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
883  | 
(* list") are identified. *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
884  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
885  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
886  | 
fun string_of_typ (Type (s, _)) = s  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
887  | 
| string_of_typ (TFree (s, _)) = s  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
888  | 
| string_of_typ (TVar ((s,_), _)) = s;  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
889  | 
|
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
890  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
891  | 
(* first_universe: returns the "first" (i.e. smallest) universe by assigning *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
892  | 
(* 'minsize' to every type for which no size is specified in *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
893  | 
(* 'sizes' *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
894  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
895  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
896  | 
fun first_universe xs sizes minsize =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
897  | 
let  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
898  | 
fun size_of_typ T =  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
899  | 
case AList.lookup (op =) sizes (string_of_typ T) of  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
900  | 
SOME n => n  | 
| 33246 | 901  | 
| NONE => minsize  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
902  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
903  | 
map (fn T => (T, size_of_typ T)) xs  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
904  | 
end;  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
905  | 
|
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
906  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
907  | 
(* next_universe: enumerates all universes (i.e. assignments of sizes to *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
908  | 
(* types), where the minimal size of a type is given by *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
909  | 
(* 'minsize', the maximal size is given by 'maxsize', and a *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
910  | 
(* type may have a fixed size given in 'sizes' *)  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
911  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
912  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
913  | 
fun next_universe xs sizes minsize maxsize =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
914  | 
let  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
915  | 
(* creates the "first" list of length 'len', where the sum of all list *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
916  | 
(* elements is 'sum', and the length of the list is 'len' *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
917  | 
fun make_first _ 0 sum =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
918  | 
if sum = 0 then  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
919  | 
SOME []  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
920  | 
else  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
921  | 
NONE  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
922  | 
| make_first max len sum =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
923  | 
if sum <= max orelse max < 0 then  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
924  | 
Option.map (fn xs' => sum :: xs') (make_first max (len-1) 0)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
925  | 
else  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
926  | 
Option.map (fn xs' => max :: xs') (make_first max (len-1) (sum-max))  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
927  | 
(* enumerates all int lists with a fixed length, where 0<=x<='max' for *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
928  | 
(* all list elements x (unless 'max'<0) *)  | 
| 46096 | 929  | 
fun next _ _ _ [] =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
930  | 
NONE  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
931  | 
| next max len sum [x] =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
932  | 
(* we've reached the last list element, so there's no shift possible *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
933  | 
make_first max (len+1) (sum+x+1) (* increment 'sum' by 1 *)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
934  | 
| next max len sum (x1::x2::xs) =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
935  | 
if x1>0 andalso (x2<max orelse max<0) then  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
936  | 
(* we can shift *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
937  | 
SOME (the (make_first max (len+1) (sum+x1-1)) @ (x2+1) :: xs)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
938  | 
else  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
939  | 
(* continue search *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
940  | 
next max (len+1) (sum+x1) (x2::xs)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
941  | 
(* only consider those types for which the size is not fixed *)  | 
| 33317 | 942  | 
val mutables = filter_out (AList.defined (op =) sizes o string_of_typ o fst) xs  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
943  | 
(* subtract 'minsize' from every size (will be added again at the end) *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
944  | 
val diffs = map (fn (_, n) => n-minsize) mutables  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
945  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
946  | 
case next (maxsize-minsize) 0 0 diffs of  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
947  | 
SOME diffs' =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
948  | 
(* merge with those types for which the size is fixed *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
949  | 
SOME (fst (fold_map (fn (T, _) => fn ds =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
950  | 
case AList.lookup (op =) sizes (string_of_typ T) of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
951  | 
(* return the fixed size *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
952  | 
SOME n => ((T, n), ds)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
953  | 
(* consume the head of 'ds', add 'minsize' *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
954  | 
| NONE => ((T, minsize + hd ds), tl ds))  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
955  | 
xs diffs'))  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
956  | 
| NONE => NONE  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
957  | 
end;  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
958  | 
|
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
959  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
960  | 
(* toTrue: converts the interpretation of a Boolean value to a propositional *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
961  | 
(* formula that is true iff the interpretation denotes "true" *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
962  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
963  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
964  | 
fun toTrue (Leaf [fm, _]) = fm  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
965  | 
  | toTrue _ = raise REFUTE ("toTrue", "interpretation does not denote a Boolean value");
 | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
966  | 
|
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
967  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
968  | 
(* toFalse: converts the interpretation of a Boolean value to a *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
969  | 
(* propositional formula that is true iff the interpretation *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
970  | 
(* denotes "false" *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
971  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
972  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
973  | 
fun toFalse (Leaf [_, fm]) = fm  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
974  | 
  | toFalse _ = raise REFUTE ("toFalse", "interpretation does not denote a Boolean value");
 | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
975  | 
|
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
976  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
977  | 
(* find_model: repeatedly calls 'interpret' with appropriate parameters, *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
978  | 
(* applies a SAT solver, and (in case a model is found) displays *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
979  | 
(* the model to the user by calling 'print_model' *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
980  | 
(* {...}     : parameters that control the translation/model generation      *)
 | 
| 
34120
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
981  | 
(* assm_ts : assumptions to be considered unless "no_assms" is specified *)  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
982  | 
(* t : term to be translated into a propositional formula *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
983  | 
(* negate : if true, find a model that makes 't' false (rather than true) *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
984  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
985  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
986  | 
fun find_model ctxt  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
987  | 
    {sizes, minsize, maxsize, maxvars, maxtime, satsolver, no_assms, expect}
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
988  | 
assm_ts t negate =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
989  | 
let  | 
| 42361 | 990  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
33054
 
dd1192a96968
fixed the "expect" mechanism of Refute in the face of timeouts
 
blanchet 
parents: 
32952 
diff
changeset
 | 
991  | 
fun check_expect outcome_code =  | 
| 
45387
 
ccffb3f9f42b
return outcome code, so that it can be picked up by Mutabelle
 
blanchet 
parents: 
44121 
diff
changeset
 | 
992  | 
if expect = "" orelse outcome_code = expect then outcome_code  | 
| 
33054
 
dd1192a96968
fixed the "expect" mechanism of Refute in the face of timeouts
 
blanchet 
parents: 
32952 
diff
changeset
 | 
993  | 
      else error ("Unexpected outcome: " ^ quote outcome_code ^ ".")
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
994  | 
fun wrapper () =  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
995  | 
let  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
996  | 
val timer = Timer.startRealTimer ()  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
997  | 
val t =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
998  | 
if no_assms then t  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
999  | 
else if negate then Logic.list_implies (assm_ts, t)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1000  | 
else Logic.mk_conjunction_list (t :: assm_ts)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1001  | 
val u = unfold_defs thy t  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1002  | 
        val _ = tracing ("Unfolded term: " ^ Syntax.string_of_term ctxt u)
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1003  | 
val axioms = collect_axioms ctxt u  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1004  | 
val types = fold (union (op =) o ground_types ctxt) (u :: axioms) []  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1005  | 
        val _ = tracing ("Ground types: "
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1006  | 
^ (if null types then "none."  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1007  | 
else commas (map (Syntax.string_of_typ ctxt) types)))  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1008  | 
(* we can only consider fragments of recursive IDTs, so we issue a *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1009  | 
(* warning if the formula contains a recursive IDT *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1010  | 
(* TODO: no warning needed for /positive/ occurrences of IDTs *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1011  | 
val maybe_spurious = Library.exists (fn  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1012  | 
Type (s, _) =>  | 
| 
58354
 
04ac60da613e
support (finite values of) codatatypes in Quickcheck
 
blanchet 
parents: 
58322 
diff
changeset
 | 
1013  | 
(case BNF_LFP_Compat.get_info thy [] s of  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1014  | 
SOME info => (* inductive datatype *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1015  | 
let  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1016  | 
val index = #index info  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1017  | 
val descr = #descr info  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1018  | 
val (_, _, constrs) = the (AList.lookup (op =) descr index)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1019  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1020  | 
(* recursive datatype? *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1021  | 
Library.exists (fn (_, ds) =>  | 
| 58120 | 1022  | 
Library.exists Old_Datatype_Aux.is_rec_type ds) constrs  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1023  | 
end  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1024  | 
| NONE => false)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1025  | 
| _ => false) types  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1026  | 
val _ =  | 
| 
59352
 
63c02d051661
tuned warnings: observe Context_Position.is_visible;
 
wenzelm 
parents: 
58893 
diff
changeset
 | 
1027  | 
if maybe_spurious andalso Context_Position.is_visible ctxt then  | 
| 
 
63c02d051661
tuned warnings: observe Context_Position.is_visible;
 
wenzelm 
parents: 
58893 
diff
changeset
 | 
1028  | 
warning "Term contains a recursive datatype; countermodel(s) may be spurious!"  | 
| 
 
63c02d051661
tuned warnings: observe Context_Position.is_visible;
 
wenzelm 
parents: 
58893 
diff
changeset
 | 
1029  | 
else ()  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1030  | 
fun find_model_loop universe =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1031  | 
let  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1032  | 
val msecs_spent = Time.toMilliseconds (Timer.checkRealTimer timer)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1033  | 
val _ = maxtime = 0 orelse msecs_spent < 1000 * maxtime  | 
| 62519 | 1034  | 
orelse raise Timeout.TIMEOUT (Time.fromMilliseconds msecs_spent)  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1035  | 
val init_model = (universe, [])  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1036  | 
            val init_args  = {maxvars = maxvars, def_eq = false, next_idx = 1,
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1037  | 
bounds = [], wellformed = True}  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1038  | 
            val _ = tracing ("Translating term (sizes: "
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1039  | 
^ commas (map (fn (_, n) => string_of_int n) universe) ^ ") ...")  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1040  | 
(* translate 'u' and all axioms *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1041  | 
val (intrs, (model, args)) = fold_map (fn t' => fn (m, a) =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1042  | 
let  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1043  | 
val (i, m', a') = interpret ctxt m a t'  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1044  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1045  | 
(* set 'def_eq' to 'true' *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1046  | 
                (i, (m', {maxvars = #maxvars a', def_eq = true,
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1047  | 
next_idx = #next_idx a', bounds = #bounds a',  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1048  | 
wellformed = #wellformed a'}))  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1049  | 
end) (u :: axioms) (init_model, init_args)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1050  | 
(* make 'u' either true or false, and make all axioms true, and *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1051  | 
(* add the well-formedness side condition *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1052  | 
val fm_u = (if negate then toFalse else toTrue) (hd intrs)  | 
| 41471 | 1053  | 
val fm_ax = Prop_Logic.all (map toTrue (tl intrs))  | 
1054  | 
val fm = Prop_Logic.all [#wellformed args, fm_ax, fm_u]  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1055  | 
val _ =  | 
| 
59352
 
63c02d051661
tuned warnings: observe Context_Position.is_visible;
 
wenzelm 
parents: 
58893 
diff
changeset
 | 
1056  | 
(if member (op =) ["cdclite"] satsolver andalso Context_Position.is_visible ctxt then  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1057  | 
                warning ("Using SAT solver " ^ quote satsolver ^
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1058  | 
"; for better performance, consider installing an \  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1059  | 
\external solver.")  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1060  | 
else ());  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1061  | 
val solver =  | 
| 56853 | 1062  | 
SAT_Solver.invoke_solver satsolver  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1063  | 
handle Option.Option =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1064  | 
                     error ("Unknown SAT solver: " ^ quote satsolver ^
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1065  | 
". Available solvers: " ^  | 
| 56853 | 1066  | 
commas (map (quote o fst) (SAT_Solver.get_solvers ())) ^ ".")  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1067  | 
in  | 
| 58843 | 1068  | 
writeln "Invoking SAT solver...";  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1069  | 
(case solver fm of  | 
| 56853 | 1070  | 
SAT_Solver.SATISFIABLE assignment =>  | 
| 58843 | 1071  | 
                (writeln ("Model found:\n" ^ print_model ctxt model
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1072  | 
(fn i => case assignment i of SOME b => b | NONE => true));  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1073  | 
if maybe_spurious then "potential" else "genuine")  | 
| 56853 | 1074  | 
| SAT_Solver.UNSATISFIABLE _ =>  | 
| 58843 | 1075  | 
(writeln "No model exists.";  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1076  | 
case next_universe universe sizes minsize maxsize of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1077  | 
SOME universe' => find_model_loop universe'  | 
| 58843 | 1078  | 
| NONE => (writeln  | 
| 
40132
 
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
 
wenzelm 
parents: 
39811 
diff
changeset
 | 
1079  | 
"Search terminated, no larger universe within the given limits.";  | 
| 
 
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
 
wenzelm 
parents: 
39811 
diff
changeset
 | 
1080  | 
"none"))  | 
| 56853 | 1081  | 
| SAT_Solver.UNKNOWN =>  | 
| 58843 | 1082  | 
(writeln "No model found.";  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1083  | 
case next_universe universe sizes minsize maxsize of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1084  | 
SOME universe' => find_model_loop universe'  | 
| 58843 | 1085  | 
| NONE => (writeln  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1086  | 
"Search terminated, no larger universe within the given limits.";  | 
| 56853 | 1087  | 
"unknown"))) handle SAT_Solver.NOT_CONFIGURED =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1088  | 
              (error ("SAT solver " ^ quote satsolver ^ " is not configured.");
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1089  | 
"unknown")  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1090  | 
end  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1091  | 
handle MAXVARS_EXCEEDED =>  | 
| 58843 | 1092  | 
            (writeln ("Search terminated, number of Boolean variables ("
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1093  | 
^ string_of_int maxvars ^ " allowed) exceeded.");  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1094  | 
"unknown")  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1095  | 
|
| 
30314
 
853778f6ef7d
Added "expect" option to Refute, like in Nitpick, that allows to write regression tests.
 
blanchet 
parents: 
30312 
diff
changeset
 | 
1096  | 
val outcome_code = find_model_loop (first_universe types sizes minsize)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1097  | 
in  | 
| 
33054
 
dd1192a96968
fixed the "expect" mechanism of Refute in the face of timeouts
 
blanchet 
parents: 
32952 
diff
changeset
 | 
1098  | 
check_expect outcome_code  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1099  | 
end  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1100  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1101  | 
(* some parameter sanity checks *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1102  | 
minsize>=1 orelse  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1103  | 
      error ("\"minsize\" is " ^ string_of_int minsize ^ ", must be at least 1");
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1104  | 
maxsize>=1 orelse  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1105  | 
      error ("\"maxsize\" is " ^ string_of_int maxsize ^ ", must be at least 1");
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1106  | 
maxsize>=minsize orelse  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1107  | 
      error ("\"maxsize\" (=" ^ string_of_int maxsize ^
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1108  | 
") is less than \"minsize\" (=" ^ string_of_int minsize ^ ").");  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1109  | 
maxvars>=0 orelse  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1110  | 
      error ("\"maxvars\" is " ^ string_of_int maxvars ^ ", must be at least 0");
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1111  | 
maxtime>=0 orelse  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1112  | 
      error ("\"maxtime\" is " ^ string_of_int maxtime ^ ", must be at least 0");
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1113  | 
(* enter loop with or without time limit *)  | 
| 58843 | 1114  | 
    writeln ("Trying to find a model that "
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1115  | 
^ (if negate then "refutes" else "satisfies") ^ ": "  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1116  | 
^ Syntax.string_of_term ctxt t);  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1117  | 
if maxtime > 0 then (  | 
| 62519 | 1118  | 
Timeout.apply (Time.fromSeconds maxtime)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1119  | 
wrapper ()  | 
| 62519 | 1120  | 
handle Timeout.TIMEOUT _ =>  | 
| 58843 | 1121  | 
        (writeln ("Search terminated, time limit (" ^
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1122  | 
string_of_int maxtime  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1123  | 
^ (if maxtime=1 then " second" else " seconds") ^ ") exceeded.");  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1124  | 
check_expect "unknown")  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1125  | 
) else wrapper ()  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1126  | 
end;  | 
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
1127  | 
|
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
1128  | 
|
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
1129  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
1130  | 
(* INTERFACE, PART 2: FINDING A MODEL *)  | 
| 14350 | 1131  | 
(* ------------------------------------------------------------------------- *)  | 
1132  | 
||
1133  | 
(* ------------------------------------------------------------------------- *)  | 
|
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
1134  | 
(* satisfy_term: calls 'find_model' to find a model that satisfies 't' *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
1135  | 
(* params : list of '(name, value)' pairs used to override default *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
1136  | 
(* parameters *)  | 
| 14350 | 1137  | 
(* ------------------------------------------------------------------------- *)  | 
1138  | 
||
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1139  | 
fun satisfy_term ctxt params assm_ts t =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1140  | 
find_model ctxt (actual_params ctxt params) assm_ts t false;  | 
| 14350 | 1141  | 
|
1142  | 
(* ------------------------------------------------------------------------- *)  | 
|
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
1143  | 
(* refute_term: calls 'find_model' to find a model that refutes 't' *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
1144  | 
(* params : list of '(name, value)' pairs used to override default *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
1145  | 
(* parameters *)  | 
| 14350 | 1146  | 
(* ------------------------------------------------------------------------- *)  | 
1147  | 
||
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1148  | 
fun refute_term ctxt params assm_ts t =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1149  | 
let  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1150  | 
(* disallow schematic type variables, since we cannot properly negate *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1151  | 
(* terms containing them (their logical meaning is that there EXISTS a *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1152  | 
(* type s.t. ...; to refute such a formula, we would have to show that *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1153  | 
(* for ALL types, not ...) *)  | 
| 29272 | 1154  | 
val _ = null (Term.add_tvars t []) orelse  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1155  | 
error "Term to be refuted contains schematic type variables"  | 
| 21556 | 1156  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1157  | 
(* existential closure over schematic variables *)  | 
| 
60924
 
610794dff23c
tuned signature, in accordance to sortBy in Scala;
 
wenzelm 
parents: 
60190 
diff
changeset
 | 
1158  | 
val vars = sort_by (fst o fst) (Term.add_vars t [])  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1159  | 
(* Term.term *)  | 
| 33246 | 1160  | 
val ex_closure = fold (fn ((x, i), T) => fn t' =>  | 
1161  | 
HOLogic.exists_const T $  | 
|
1162  | 
Abs (x, T, abstract_over (Var ((x, i), T), t'))) vars t  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1163  | 
(* Note: If 't' is of type 'propT' (rather than 'boolT'), applying *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1164  | 
(* 'HOLogic.exists_const' is not type-correct. However, this is not *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1165  | 
(* really a problem as long as 'find_model' still interprets the *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1166  | 
(* resulting term correctly, without checking its type. *)  | 
| 21556 | 1167  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1168  | 
(* replace outermost universally quantified variables by Free's: *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1169  | 
(* refuting a term with Free's is generally faster than refuting a *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1170  | 
(* term with (nested) quantifiers, because quantifiers are expanded, *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1171  | 
(* while the SAT solver searches for an interpretation for Free's. *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1172  | 
(* Also we get more information back that way, namely an *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1173  | 
(* interpretation which includes values for the (formerly) *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1174  | 
(* quantified variables. *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1175  | 
(* maps !!x1...xn. !xk...xm. t to t *)  | 
| 56245 | 1176  | 
    fun strip_all_body (Const (@{const_name Pure.all}, _) $ Abs (_, _, t)) =
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1177  | 
strip_all_body t  | 
| 
29820
 
07f53494cf20
Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
 
blanchet 
parents: 
29802 
diff
changeset
 | 
1178  | 
      | strip_all_body (Const (@{const_name Trueprop}, _) $ t) =
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1179  | 
strip_all_body t  | 
| 
29820
 
07f53494cf20
Rearrange Refute/SAT theory dependencies so as to use even more antiquotations in refute.ML +
 
blanchet 
parents: 
29802 
diff
changeset
 | 
1180  | 
      | strip_all_body (Const (@{const_name All}, _) $ Abs (_, _, t)) =
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1181  | 
strip_all_body t  | 
| 
29802
 
d201a838d0f7
Make some Refute functions public so I can use them in Nitpick,
 
blanchet 
parents: 
29288 
diff
changeset
 | 
1182  | 
| strip_all_body t = t  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1183  | 
(* maps !!x1...xn. !xk...xm. t to [x1, ..., xn, xk, ..., xm] *)  | 
| 56245 | 1184  | 
    fun strip_all_vars (Const (@{const_name Pure.all}, _) $ Abs (a, T, t)) =
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1185  | 
(a, T) :: strip_all_vars t  | 
| 
29802
 
d201a838d0f7
Make some Refute functions public so I can use them in Nitpick,
 
blanchet 
parents: 
29288 
diff
changeset
 | 
1186  | 
      | strip_all_vars (Const (@{const_name Trueprop}, _) $ t) =
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1187  | 
strip_all_vars t  | 
| 
29802
 
d201a838d0f7
Make some Refute functions public so I can use them in Nitpick,
 
blanchet 
parents: 
29288 
diff
changeset
 | 
1188  | 
      | strip_all_vars (Const (@{const_name All}, _) $ Abs (a, T, t)) =
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1189  | 
(a, T) :: strip_all_vars t  | 
| 46096 | 1190  | 
| strip_all_vars _ = [] : (string * typ) list  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1191  | 
val strip_t = strip_all_body ex_closure  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1192  | 
val frees = Term.rename_wrt_term strip_t (strip_all_vars ex_closure)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1193  | 
val subst_t = Term.subst_bounds (map Free frees, strip_t)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1194  | 
in  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1195  | 
find_model ctxt (actual_params ctxt params) assm_ts subst_t true  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1196  | 
end;  | 
| 14350 | 1197  | 
|
1198  | 
(* ------------------------------------------------------------------------- *)  | 
|
| 
32857
 
394d37f19e0a
Refute.refute_goal: canonical goal addresses from 1 (renamed from refute_subgoal to clarify change in semantics);
 
wenzelm 
parents: 
31986 
diff
changeset
 | 
1199  | 
(* refute_goal *)  | 
| 14350 | 1200  | 
(* ------------------------------------------------------------------------- *)  | 
1201  | 
||
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1202  | 
fun refute_goal ctxt params th i =  | 
| 
34120
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
1203  | 
let  | 
| 59582 | 1204  | 
val t = th |> Thm.prop_of  | 
| 
34120
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
1205  | 
in  | 
| 
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
1206  | 
if Logic.count_prems t = 0 then  | 
| 58843 | 1207  | 
(writeln "No subgoal!"; "none")  | 
| 
34120
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
1208  | 
else  | 
| 
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
1209  | 
let  | 
| 59582 | 1210  | 
val assms = map Thm.term_of (Assumption.all_assms_of ctxt)  | 
| 
34120
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
1211  | 
val (t, frees) = Logic.goal_params t i  | 
| 
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
1212  | 
in  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1213  | 
refute_term ctxt params assms (subst_bounds (frees, t))  | 
| 
34120
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
1214  | 
end  | 
| 
 
f9920a3ddf50
added "no_assms" option to Refute, and include structured proof assumptions by default;
 
blanchet 
parents: 
34017 
diff
changeset
 | 
1215  | 
end  | 
| 14350 | 1216  | 
|
1217  | 
||
1218  | 
(* ------------------------------------------------------------------------- *)  | 
|
| 15292 | 1219  | 
(* INTERPRETERS: Auxiliary Functions *)  | 
| 14350 | 1220  | 
(* ------------------------------------------------------------------------- *)  | 
1221  | 
||
1222  | 
(* ------------------------------------------------------------------------- *)  | 
|
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1223  | 
(* make_constants: returns all interpretations for type 'T' that consist of *)  | 
| 
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1224  | 
(* unit vectors with 'True'/'False' only (no Boolean *)  | 
| 
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1225  | 
(* variables) *)  | 
| 14350 | 1226  | 
(* ------------------------------------------------------------------------- *)  | 
1227  | 
||
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1228  | 
fun make_constants ctxt model T =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1229  | 
let  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1230  | 
(* returns a list with all unit vectors of length n *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1231  | 
fun unit_vectors n =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1232  | 
let  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1233  | 
(* returns the k-th unit vector of length n *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1234  | 
fun unit_vector (k, n) =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1235  | 
Leaf ((replicate (k-1) False) @ (True :: (replicate (n-k) False)))  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1236  | 
fun unit_vectors_loop k =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1237  | 
if k>n then [] else unit_vector (k,n) :: unit_vectors_loop (k+1)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1238  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1239  | 
unit_vectors_loop 1  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1240  | 
end  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1241  | 
(* returns a list of lists, each one consisting of n (possibly *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1242  | 
(* identical) elements from 'xs' *)  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1243  | 
fun pick_all 1 xs = map single xs  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1244  | 
| pick_all n xs =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1245  | 
let val rec_pick = pick_all (n - 1) xs in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1246  | 
maps (fn x => map (cons x) rec_pick) xs  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1247  | 
end  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1248  | 
(* returns all constant interpretations that have the same tree *)  | 
| 
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1249  | 
(* structure as the interpretation argument *)  | 
| 
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1250  | 
fun make_constants_intr (Leaf xs) = unit_vectors (length xs)  | 
| 
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1251  | 
| make_constants_intr (Node xs) = map Node (pick_all (length xs)  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1252  | 
(make_constants_intr (hd xs)))  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1253  | 
(* obtain the interpretation for a variable of type 'T' *)  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1254  | 
    val (i, _, _) = interpret ctxt model {maxvars=0, def_eq=false, next_idx=1,
 | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1255  | 
      bounds=[], wellformed=True} (Free ("dummy", T))
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1256  | 
in  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1257  | 
make_constants_intr i  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1258  | 
end;  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1259  | 
|
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1260  | 
(* ------------------------------------------------------------------------- *)  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1261  | 
(* size_of_type: returns the number of elements in a type 'T' (i.e. 'length *)  | 
| 
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1262  | 
(* (make_constants T)', but implemented more efficiently) *)  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1263  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1264  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1265  | 
(* returns 0 for an empty ground type or a function type with empty *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1266  | 
(* codomain, but fails for a function type with empty domain -- *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1267  | 
(* admissibility of datatype constructor argument types (see "Inductive *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1268  | 
(* datatypes in HOL - lessons learned ...", S. Berghofer, M. Wenzel, *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1269  | 
(* TPHOLs 99) ensures that recursive, possibly empty, datatype fragments *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1270  | 
(* never occur as the domain of a function type that is the type of a *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1271  | 
(* constructor argument *)  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1272  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1273  | 
fun size_of_type ctxt model T =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1274  | 
let  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1275  | 
(* returns the number of elements that have the same tree structure as a *)  | 
| 
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1276  | 
(* given interpretation *)  | 
| 
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1277  | 
fun size_of_intr (Leaf xs) = length xs  | 
| 
39047
 
cdff476ba39e
use existing Integer.pow, despite its slightly odd argument order;
 
wenzelm 
parents: 
39046 
diff
changeset
 | 
1278  | 
| size_of_intr (Node xs) = Integer.pow (length xs) (size_of_intr (hd xs))  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1279  | 
(* obtain the interpretation for a variable of type 'T' *)  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1280  | 
    val (i, _, _) = interpret ctxt model {maxvars=0, def_eq=false, next_idx=1,
 | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1281  | 
      bounds=[], wellformed=True} (Free ("dummy", T))
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1282  | 
in  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1283  | 
size_of_intr i  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1284  | 
end;  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1285  | 
|
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1286  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1287  | 
(* TT/FF: interpretations that denote "true" or "false", respectively *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1288  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1289  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1290  | 
val TT = Leaf [True, False];  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1291  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1292  | 
val FF = Leaf [False, True];  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1293  | 
|
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1294  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1295  | 
(* make_equality: returns an interpretation that denotes (extensional) *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1296  | 
(* equality of two interpretations *)  | 
| 
15547
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1297  | 
(* - two interpretations are 'equal' iff they are both defined and denote *)  | 
| 
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1298  | 
(* the same value *)  | 
| 
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1299  | 
(* - two interpretations are 'not_equal' iff they are both defined at least *)  | 
| 
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1300  | 
(* partially, and a defined part denotes different values *)  | 
| 
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1301  | 
(* - a completely undefined interpretation is neither 'equal' nor *)  | 
| 
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1302  | 
(* 'not_equal' to another interpretation *)  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1303  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1304  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1305  | 
(* We could in principle represent '=' on a type T by a particular *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1306  | 
(* interpretation. However, the size of that interpretation is quadratic *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1307  | 
(* in the size of T. Therefore comparing the interpretations 'i1' and *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1308  | 
(* 'i2' directly is more efficient than constructing the interpretation *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1309  | 
(* for equality on T first, and "applying" this interpretation to 'i1' *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1310  | 
(* and 'i2' in the usual way (cf. 'interpretation_apply') then. *)  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1311  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1312  | 
fun make_equality (i1, i2) =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1313  | 
let  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1314  | 
fun equal (i1, i2) =  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1315  | 
(case i1 of  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1316  | 
Leaf xs =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1317  | 
(case i2 of  | 
| 41471 | 1318  | 
Leaf ys => Prop_Logic.dot_product (xs, ys) (* defined and equal *)  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1319  | 
          | Node _  => raise REFUTE ("make_equality",
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1320  | 
"second interpretation is higher"))  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1321  | 
| Node xs =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1322  | 
(case i2 of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1323  | 
            Leaf _  => raise REFUTE ("make_equality",
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1324  | 
"first interpretation is higher")  | 
| 41471 | 1325  | 
| Node ys => Prop_Logic.all (map equal (xs ~~ ys))))  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1326  | 
fun not_equal (i1, i2) =  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1327  | 
(case i1 of  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1328  | 
Leaf xs =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1329  | 
(case i2 of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1330  | 
(* defined and not equal *)  | 
| 41471 | 1331  | 
Leaf ys => Prop_Logic.all ((Prop_Logic.exists xs)  | 
1332  | 
:: (Prop_Logic.exists ys)  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1333  | 
:: (map (fn (x,y) => SOr (SNot x, SNot y)) (xs ~~ ys)))  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1334  | 
          | Node _  => raise REFUTE ("make_equality",
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1335  | 
"second interpretation is higher"))  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1336  | 
| Node xs =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1337  | 
(case i2 of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1338  | 
            Leaf _  => raise REFUTE ("make_equality",
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1339  | 
"first interpretation is higher")  | 
| 41471 | 1340  | 
| Node ys => Prop_Logic.exists (map not_equal (xs ~~ ys))))  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1341  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1342  | 
(* a value may be undefined; therefore 'not_equal' is not just the *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1343  | 
(* negation of 'equal' *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1344  | 
Leaf [equal (i1, i2), not_equal (i1, i2)]  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1345  | 
end;  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1346  | 
|
| 15292 | 1347  | 
(* ------------------------------------------------------------------------- *)  | 
| 
15547
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1348  | 
(* make_def_equality: returns an interpretation that denotes (extensional) *)  | 
| 
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1349  | 
(* equality of two interpretations *)  | 
| 
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1350  | 
(* This function treats undefined/partially defined interpretations *)  | 
| 
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1351  | 
(* different from 'make_equality': two undefined interpretations are *)  | 
| 
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1352  | 
(* considered equal, while a defined interpretation is considered not equal *)  | 
| 
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1353  | 
(* to an undefined interpretation. *)  | 
| 
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1354  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1355  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1356  | 
fun make_def_equality (i1, i2) =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1357  | 
let  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1358  | 
fun equal (i1, i2) =  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1359  | 
(case i1 of  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1360  | 
Leaf xs =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1361  | 
(case i2 of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1362  | 
(* defined and equal, or both undefined *)  | 
| 41471 | 1363  | 
Leaf ys => SOr (Prop_Logic.dot_product (xs, ys),  | 
1364  | 
SAnd (Prop_Logic.all (map SNot xs), Prop_Logic.all (map SNot ys)))  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1365  | 
          | Node _  => raise REFUTE ("make_def_equality",
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1366  | 
"second interpretation is higher"))  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1367  | 
| Node xs =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1368  | 
(case i2 of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1369  | 
            Leaf _  => raise REFUTE ("make_def_equality",
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1370  | 
"first interpretation is higher")  | 
| 41471 | 1371  | 
| Node ys => Prop_Logic.all (map equal (xs ~~ ys))))  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1372  | 
val eq = equal (i1, i2)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1373  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1374  | 
Leaf [eq, SNot eq]  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1375  | 
end;  | 
| 
15547
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1376  | 
|
| 
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1377  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1378  | 
(* interpretation_apply: returns an interpretation that denotes the result *)  | 
| 22092 | 1379  | 
(* of applying the function denoted by 'i1' to the *)  | 
| 
15547
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1380  | 
(* argument denoted by 'i2' *)  | 
| 
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1381  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1382  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1383  | 
fun interpretation_apply (i1, i2) =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1384  | 
let  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1385  | 
fun interpretation_disjunction (tr1,tr2) =  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1386  | 
tree_map (fn (xs,ys) => map (fn (x,y) => SOr(x,y)) (xs ~~ ys))  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1387  | 
(tree_pair (tr1,tr2))  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1388  | 
fun prop_formula_times_interpretation (fm,tr) =  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1389  | 
tree_map (map (fn x => SAnd (fm,x))) tr  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1390  | 
fun prop_formula_list_dot_product_interpretation_list ([fm],[tr]) =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1391  | 
prop_formula_times_interpretation (fm,tr)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1392  | 
| prop_formula_list_dot_product_interpretation_list (fm::fms,tr::trees) =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1393  | 
interpretation_disjunction (prop_formula_times_interpretation (fm,tr),  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1394  | 
prop_formula_list_dot_product_interpretation_list (fms,trees))  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1395  | 
| prop_formula_list_dot_product_interpretation_list (_,_) =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1396  | 
          raise REFUTE ("interpretation_apply", "empty list (in dot product)")
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1397  | 
(* returns a list of lists, each one consisting of one element from each *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1398  | 
(* element of 'xss' *)  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1399  | 
fun pick_all [xs] = map single xs  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1400  | 
| pick_all (xs::xss) =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1401  | 
let val rec_pick = pick_all xss in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1402  | 
maps (fn x => map (cons x) rec_pick) xs  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1403  | 
end  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1404  | 
      | pick_all _ = raise REFUTE ("interpretation_apply", "empty list (in pick_all)")
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1405  | 
fun interpretation_to_prop_formula_list (Leaf xs) = xs  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1406  | 
| interpretation_to_prop_formula_list (Node trees) =  | 
| 41471 | 1407  | 
map Prop_Logic.all (pick_all  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1408  | 
(map interpretation_to_prop_formula_list trees))  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1409  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1410  | 
case i1 of  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1411  | 
Leaf _ =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1412  | 
        raise REFUTE ("interpretation_apply", "first interpretation is a leaf")
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1413  | 
| Node xs =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1414  | 
prop_formula_list_dot_product_interpretation_list  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1415  | 
(interpretation_to_prop_formula_list i2, xs)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1416  | 
end;  | 
| 
15547
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1417  | 
|
| 
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1418  | 
(* ------------------------------------------------------------------------- *)  | 
| 15292 | 1419  | 
(* eta_expand: eta-expands a term 't' by adding 'i' lambda abstractions *)  | 
1420  | 
(* ------------------------------------------------------------------------- *)  | 
|
1421  | 
||
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1422  | 
fun eta_expand t i =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1423  | 
let  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1424  | 
val Ts = Term.binder_types (Term.fastype_of t)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1425  | 
val t' = Term.incr_boundvars i t  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1426  | 
in  | 
| 33339 | 1427  | 
    fold_rev (fn T => fn term => Abs ("<eta_expand>", T, term))
 | 
1428  | 
(List.take (Ts, i))  | 
|
1429  | 
(Term.list_comb (t', map Bound (i-1 downto 0)))  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1430  | 
end;  | 
| 15292 | 1431  | 
|
| 
15335
 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 
webertj 
parents: 
15334 
diff
changeset
 | 
1432  | 
(* ------------------------------------------------------------------------- *)  | 
| 
15547
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1433  | 
(* size_of_dtyp: the size of (an initial fragment of) an inductive data type *)  | 
| 
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1434  | 
(* is the sum (over its constructors) of the product (over *)  | 
| 
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1435  | 
(* their arguments) of the size of the argument types *)  | 
| 
15335
 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 
webertj 
parents: 
15334 
diff
changeset
 | 
1436  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 
webertj 
parents: 
15334 
diff
changeset
 | 
1437  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1438  | 
fun size_of_dtyp ctxt typ_sizes descr typ_assoc constructors =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1439  | 
Integer.sum (map (fn (_, dtyps) =>  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1440  | 
Integer.prod (map (size_of_type ctxt (typ_sizes, []) o  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1441  | 
(typ_of_dtyp descr typ_assoc)) dtyps))  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1442  | 
constructors);  | 
| 
15335
 
f81e6e24351f
minor code refactoring (typ_of_dtyp, size_of_dtyp)
 
webertj 
parents: 
15334 
diff
changeset
 | 
1443  | 
|
| 15292 | 1444  | 
|
1445  | 
(* ------------------------------------------------------------------------- *)  | 
|
1446  | 
(* INTERPRETERS: Actual Interpreters *)  | 
|
1447  | 
(* ------------------------------------------------------------------------- *)  | 
|
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1448  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1449  | 
(* simply typed lambda calculus: Isabelle's basic term syntax, with type *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1450  | 
(* variables, function types, and propT *)  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1451  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1452  | 
fun stlc_interpreter ctxt model args t =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1453  | 
let  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1454  | 
val (typs, terms) = model  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1455  | 
    val {maxvars, def_eq, next_idx, bounds, wellformed} = args
 | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1456  | 
fun interpret_groundterm T =  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1457  | 
let  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1458  | 
fun interpret_groundtype () =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1459  | 
let  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1460  | 
(* the model must specify a size for ground types *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1461  | 
val size =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1462  | 
if T = Term.propT then 2  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1463  | 
else the (AList.lookup (op =) typs T)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1464  | 
val next = next_idx + size  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1465  | 
(* check if 'maxvars' is large enough *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1466  | 
val _ = (if next - 1 > maxvars andalso maxvars > 0 then  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1467  | 
raise MAXVARS_EXCEEDED else ())  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1468  | 
val fms = map BoolVar (next_idx upto (next_idx + size - 1))  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1469  | 
val intr = Leaf fms  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1470  | 
fun one_of_two_false [] = True  | 
| 41471 | 1471  | 
| one_of_two_false (x::xs) = SAnd (Prop_Logic.all (map (fn x' =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1472  | 
SOr (SNot x, SNot x')) xs), one_of_two_false xs)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1473  | 
val wf = one_of_two_false fms  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1474  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1475  | 
(* extend the model, increase 'next_idx', add well-formedness *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1476  | 
(* condition *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1477  | 
            SOME (intr, (typs, (t, intr)::terms), {maxvars = maxvars,
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1478  | 
def_eq = def_eq, next_idx = next, bounds = bounds,  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1479  | 
wellformed = SAnd (wellformed, wf)})  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1480  | 
end  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1481  | 
in  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1482  | 
case T of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1483  | 
          Type ("fun", [T1, T2]) =>
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1484  | 
let  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1485  | 
(* we create 'size_of_type ... T1' different copies of the *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1486  | 
(* interpretation for 'T2', which are then combined into a single *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1487  | 
(* new interpretation *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1488  | 
(* make fresh copies, with different variable indices *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1489  | 
(* 'idx': next variable index *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1490  | 
(* 'n' : number of copies *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1491  | 
fun make_copies idx 0 = (idx, [], True)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1492  | 
| make_copies idx n =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1493  | 
let  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1494  | 
val (copy, _, new_args) = interpret ctxt (typs, [])  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1495  | 
                        {maxvars = maxvars, def_eq = false, next_idx = idx,
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1496  | 
                        bounds = [], wellformed = True} (Free ("dummy", T2))
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1497  | 
val (idx', copies, wf') = make_copies (#next_idx new_args) (n-1)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1498  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1499  | 
(idx', copy :: copies, SAnd (#wellformed new_args, wf'))  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1500  | 
end  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1501  | 
val (next, copies, wf) = make_copies next_idx  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1502  | 
(size_of_type ctxt model T1)  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1503  | 
(* combine copies into a single interpretation *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1504  | 
val intr = Node copies  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1505  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1506  | 
(* extend the model, increase 'next_idx', add well-formedness *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1507  | 
(* condition *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1508  | 
              SOME (intr, (typs, (t, intr)::terms), {maxvars = maxvars,
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1509  | 
def_eq = def_eq, next_idx = next, bounds = bounds,  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1510  | 
wellformed = SAnd (wellformed, wf)})  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1511  | 
end  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1512  | 
| Type _ => interpret_groundtype ()  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1513  | 
| TFree _ => interpret_groundtype ()  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1514  | 
| TVar _ => interpret_groundtype ()  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1515  | 
end  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1516  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1517  | 
case AList.lookup (op =) terms t of  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1518  | 
SOME intr =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1519  | 
(* return an existing interpretation *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1520  | 
SOME (intr, model, args)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1521  | 
| NONE =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1522  | 
(case t of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1523  | 
Const (_, T) => interpret_groundterm T  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1524  | 
| Free (_, T) => interpret_groundterm T  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1525  | 
| Var (_, T) => interpret_groundterm T  | 
| 42364 | 1526  | 
| Bound i => SOME (nth (#bounds args) i, model, args)  | 
| 46096 | 1527  | 
| Abs (_, T, body) =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1528  | 
let  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1529  | 
(* create all constants of type 'T' *)  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1530  | 
val constants = make_constants ctxt model T  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1531  | 
(* interpret the 'body' separately for each constant *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1532  | 
val (bodies, (model', args')) = fold_map  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1533  | 
(fn c => fn (m, a) =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1534  | 
let  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1535  | 
(* add 'c' to 'bounds' *)  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1536  | 
                    val (i', m', a') = interpret ctxt m {maxvars = #maxvars a,
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1537  | 
def_eq = #def_eq a, next_idx = #next_idx a,  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1538  | 
bounds = (c :: #bounds a), wellformed = #wellformed a} body  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1539  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1540  | 
(* keep the new model m' and 'next_idx' and 'wellformed', *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1541  | 
(* but use old 'bounds' *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1542  | 
                    (i', (m', {maxvars = maxvars, def_eq = def_eq,
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1543  | 
next_idx = #next_idx a', bounds = bounds,  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1544  | 
wellformed = #wellformed a'}))  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1545  | 
end)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1546  | 
constants (model, args)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1547  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1548  | 
SOME (Node bodies, model', args')  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1549  | 
end  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1550  | 
| t1 $ t2 =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1551  | 
let  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1552  | 
(* interpret 't1' and 't2' separately *)  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1553  | 
val (intr1, model1, args1) = interpret ctxt model args t1  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1554  | 
val (intr2, model2, args2) = interpret ctxt model1 args1 t2  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1555  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1556  | 
SOME (interpretation_apply (intr1, intr2), model2, args2)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1557  | 
end)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1558  | 
end;  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1559  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1560  | 
fun Pure_interpreter ctxt model args t =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1561  | 
case t of  | 
| 56245 | 1562  | 
    Const (@{const_name Pure.all}, _) $ t1 =>
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1563  | 
let  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1564  | 
val (i, m, a) = interpret ctxt model args t1  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1565  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1566  | 
case i of  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1567  | 
Node xs =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1568  | 
(* 3-valued logic *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1569  | 
let  | 
| 41471 | 1570  | 
val fmTrue = Prop_Logic.all (map toTrue xs)  | 
1571  | 
val fmFalse = Prop_Logic.exists (map toFalse xs)  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1572  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1573  | 
SOME (Leaf [fmTrue, fmFalse], m, a)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1574  | 
end  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1575  | 
| _ =>  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1576  | 
          raise REFUTE ("Pure_interpreter",
 | 
| 56245 | 1577  | 
"\"Pure.all\" is followed by a non-function")  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1578  | 
end  | 
| 56245 | 1579  | 
  | Const (@{const_name Pure.all}, _) =>
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1580  | 
SOME (interpret ctxt model args (eta_expand t 1))  | 
| 56245 | 1581  | 
  | Const (@{const_name Pure.eq}, _) $ t1 $ t2 =>
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1582  | 
let  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1583  | 
val (i1, m1, a1) = interpret ctxt model args t1  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1584  | 
val (i2, m2, a2) = interpret ctxt m1 a1 t2  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1585  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1586  | 
(* we use either 'make_def_equality' or 'make_equality' *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1587  | 
SOME ((if #def_eq args then make_def_equality else make_equality)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1588  | 
(i1, i2), m2, a2)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1589  | 
end  | 
| 56245 | 1590  | 
  | Const (@{const_name Pure.eq}, _) $ _ =>
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1591  | 
SOME (interpret ctxt model args (eta_expand t 1))  | 
| 56245 | 1592  | 
  | Const (@{const_name Pure.eq}, _) =>
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1593  | 
SOME (interpret ctxt model args (eta_expand t 2))  | 
| 56245 | 1594  | 
  | Const (@{const_name Pure.imp}, _) $ t1 $ t2 =>
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1595  | 
(* 3-valued logic *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1596  | 
let  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1597  | 
val (i1, m1, a1) = interpret ctxt model args t1  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1598  | 
val (i2, m2, a2) = interpret ctxt m1 a1 t2  | 
| 41471 | 1599  | 
val fmTrue = Prop_Logic.SOr (toFalse i1, toTrue i2)  | 
1600  | 
val fmFalse = Prop_Logic.SAnd (toTrue i1, toFalse i2)  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1601  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1602  | 
SOME (Leaf [fmTrue, fmFalse], m2, a2)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1603  | 
end  | 
| 56245 | 1604  | 
  | Const (@{const_name Pure.imp}, _) $ _ =>
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1605  | 
SOME (interpret ctxt model args (eta_expand t 1))  | 
| 56245 | 1606  | 
  | Const (@{const_name Pure.imp}, _) =>
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1607  | 
SOME (interpret ctxt model args (eta_expand t 2))  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1608  | 
| _ => NONE;  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1609  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1610  | 
fun HOLogic_interpreter ctxt model args t =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1611  | 
(* Providing interpretations directly is more efficient than unfolding the *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1612  | 
(* logical constants. In HOL however, logical constants can themselves be *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1613  | 
(* arguments. They are then translated using eta-expansion. *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1614  | 
case t of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1615  | 
    Const (@{const_name Trueprop}, _) =>
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1616  | 
SOME (Node [TT, FF], model, args)  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1617  | 
  | Const (@{const_name Not}, _) =>
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1618  | 
SOME (Node [FF, TT], model, args)  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1619  | 
(* redundant, since 'True' is also an IDT constructor *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1620  | 
  | Const (@{const_name True}, _) =>
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1621  | 
SOME (TT, model, args)  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1622  | 
(* redundant, since 'False' is also an IDT constructor *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1623  | 
  | Const (@{const_name False}, _) =>
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1624  | 
SOME (FF, model, args)  | 
| 56245 | 1625  | 
  | Const (@{const_name All}, _) $ t1 =>  (* similar to "Pure.all" *)
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1626  | 
let  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1627  | 
val (i, m, a) = interpret ctxt model args t1  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1628  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1629  | 
case i of  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1630  | 
Node xs =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1631  | 
(* 3-valued logic *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1632  | 
let  | 
| 41471 | 1633  | 
val fmTrue = Prop_Logic.all (map toTrue xs)  | 
1634  | 
val fmFalse = Prop_Logic.exists (map toFalse xs)  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1635  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1636  | 
SOME (Leaf [fmTrue, fmFalse], m, a)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1637  | 
end  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1638  | 
| _ =>  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1639  | 
          raise REFUTE ("HOLogic_interpreter",
 | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1640  | 
"\"All\" is followed by a non-function")  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1641  | 
end  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1642  | 
  | Const (@{const_name All}, _) =>
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1643  | 
SOME (interpret ctxt model args (eta_expand t 1))  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1644  | 
  | Const (@{const_name Ex}, _) $ t1 =>
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1645  | 
let  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1646  | 
val (i, m, a) = interpret ctxt model args t1  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1647  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1648  | 
case i of  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1649  | 
Node xs =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1650  | 
(* 3-valued logic *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1651  | 
let  | 
| 41471 | 1652  | 
val fmTrue = Prop_Logic.exists (map toTrue xs)  | 
1653  | 
val fmFalse = Prop_Logic.all (map toFalse xs)  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1654  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1655  | 
SOME (Leaf [fmTrue, fmFalse], m, a)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1656  | 
end  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1657  | 
| _ =>  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1658  | 
          raise REFUTE ("HOLogic_interpreter",
 | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1659  | 
"\"Ex\" is followed by a non-function")  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1660  | 
end  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1661  | 
  | Const (@{const_name Ex}, _) =>
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1662  | 
SOME (interpret ctxt model args (eta_expand t 1))  | 
| 56245 | 1663  | 
  | Const (@{const_name HOL.eq}, _) $ t1 $ t2 =>  (* similar to Pure.eq *)
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1664  | 
let  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1665  | 
val (i1, m1, a1) = interpret ctxt model args t1  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1666  | 
val (i2, m2, a2) = interpret ctxt m1 a1 t2  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1667  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1668  | 
SOME (make_equality (i1, i2), m2, a2)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1669  | 
end  | 
| 46096 | 1670  | 
  | Const (@{const_name HOL.eq}, _) $ _ =>
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1671  | 
SOME (interpret ctxt model args (eta_expand t 1))  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1672  | 
  | Const (@{const_name HOL.eq}, _) =>
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1673  | 
SOME (interpret ctxt model args (eta_expand t 2))  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1674  | 
  | Const (@{const_name HOL.conj}, _) $ t1 $ t2 =>
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1675  | 
(* 3-valued logic *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1676  | 
let  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1677  | 
val (i1, m1, a1) = interpret ctxt model args t1  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1678  | 
val (i2, m2, a2) = interpret ctxt m1 a1 t2  | 
| 41471 | 1679  | 
val fmTrue = Prop_Logic.SAnd (toTrue i1, toTrue i2)  | 
1680  | 
val fmFalse = Prop_Logic.SOr (toFalse i1, toFalse i2)  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1681  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1682  | 
SOME (Leaf [fmTrue, fmFalse], m2, a2)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1683  | 
end  | 
| 46096 | 1684  | 
  | Const (@{const_name HOL.conj}, _) $ _ =>
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1685  | 
SOME (interpret ctxt model args (eta_expand t 1))  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1686  | 
  | Const (@{const_name HOL.conj}, _) =>
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1687  | 
SOME (interpret ctxt model args (eta_expand t 2))  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1688  | 
(* this would make "undef" propagate, even for formulae like *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1689  | 
(* "False & undef": *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1690  | 
(* SOME (Node [Node [TT, FF], Node [FF, FF]], model, args) *)  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1691  | 
  | Const (@{const_name HOL.disj}, _) $ t1 $ t2 =>
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1692  | 
(* 3-valued logic *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1693  | 
let  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1694  | 
val (i1, m1, a1) = interpret ctxt model args t1  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1695  | 
val (i2, m2, a2) = interpret ctxt m1 a1 t2  | 
| 41471 | 1696  | 
val fmTrue = Prop_Logic.SOr (toTrue i1, toTrue i2)  | 
1697  | 
val fmFalse = Prop_Logic.SAnd (toFalse i1, toFalse i2)  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1698  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1699  | 
SOME (Leaf [fmTrue, fmFalse], m2, a2)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1700  | 
end  | 
| 46096 | 1701  | 
  | Const (@{const_name HOL.disj}, _) $ _ =>
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1702  | 
SOME (interpret ctxt model args (eta_expand t 1))  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1703  | 
  | Const (@{const_name HOL.disj}, _) =>
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1704  | 
SOME (interpret ctxt model args (eta_expand t 2))  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1705  | 
(* this would make "undef" propagate, even for formulae like *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1706  | 
(* "True | undef": *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1707  | 
(* SOME (Node [Node [TT, TT], Node [TT, FF]], model, args) *)  | 
| 56245 | 1708  | 
  | Const (@{const_name HOL.implies}, _) $ t1 $ t2 =>  (* similar to Pure.imp *)
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1709  | 
(* 3-valued logic *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1710  | 
let  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1711  | 
val (i1, m1, a1) = interpret ctxt model args t1  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1712  | 
val (i2, m2, a2) = interpret ctxt m1 a1 t2  | 
| 41471 | 1713  | 
val fmTrue = Prop_Logic.SOr (toFalse i1, toTrue i2)  | 
1714  | 
val fmFalse = Prop_Logic.SAnd (toTrue i1, toFalse i2)  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1715  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1716  | 
SOME (Leaf [fmTrue, fmFalse], m2, a2)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1717  | 
end  | 
| 46096 | 1718  | 
  | Const (@{const_name HOL.implies}, _) $ _ =>
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1719  | 
SOME (interpret ctxt model args (eta_expand t 1))  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1720  | 
  | Const (@{const_name HOL.implies}, _) =>
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1721  | 
SOME (interpret ctxt model args (eta_expand t 2))  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1722  | 
(* this would make "undef" propagate, even for formulae like *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1723  | 
(* "False --> undef": *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1724  | 
(* SOME (Node [Node [TT, FF], Node [TT, TT]], model, args) *)  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1725  | 
| _ => NONE;  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
1726  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1727  | 
(* interprets variables and constants whose type is an IDT (this is *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1728  | 
(* relatively easy and merely requires us to compute the size of the IDT); *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1729  | 
(* constructors of IDTs however are properly interpreted by *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1730  | 
(* 'IDT_constructor_interpreter' *)  | 
| 
15547
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1731  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1732  | 
fun IDT_interpreter ctxt model args t =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1733  | 
let  | 
| 42361 | 1734  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1735  | 
val (typs, terms) = model  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1736  | 
fun interpret_term (Type (s, Ts)) =  | 
| 
58354
 
04ac60da613e
support (finite values of) codatatypes in Quickcheck
 
blanchet 
parents: 
58322 
diff
changeset
 | 
1737  | 
(case BNF_LFP_Compat.get_info thy [] s of  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1738  | 
SOME info => (* inductive datatype *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1739  | 
let  | 
| 
55507
 
5f27fb2110e0
removed odd comments -- inferred types are shown by Prover IDE;
 
wenzelm 
parents: 
55411 
diff
changeset
 | 
1740  | 
(* only recursive IDTs have an associated depth *)  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1741  | 
val depth = AList.lookup (op =) typs (Type (s, Ts))  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1742  | 
(* sanity check: depth must be at least 0 *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1743  | 
val _ =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1744  | 
(case depth of SOME n =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1745  | 
if n < 0 then  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1746  | 
                      raise REFUTE ("IDT_interpreter", "negative depth")
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1747  | 
else ()  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1748  | 
| _ => ())  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1749  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1750  | 
(* termination condition to avoid infinite recursion *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1751  | 
if depth = (SOME 0) then  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1752  | 
(* return a leaf of size 0 *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1753  | 
SOME (Leaf [], model, args)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1754  | 
else  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1755  | 
let  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1756  | 
val index = #index info  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1757  | 
val descr = #descr info  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1758  | 
val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1759  | 
val typ_assoc = dtyps ~~ Ts  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1760  | 
(* sanity check: every element in 'dtyps' must be a 'DtTFree' *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1761  | 
val _ =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1762  | 
if Library.exists (fn d =>  | 
| 58120 | 1763  | 
case d of Old_Datatype_Aux.DtTFree _ => false | _ => true) dtyps  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1764  | 
then  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1765  | 
                        raise REFUTE ("IDT_interpreter",
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1766  | 
"datatype argument (for type "  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1767  | 
^ Syntax.string_of_typ ctxt (Type (s, Ts))  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1768  | 
^ ") is not a variable")  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1769  | 
else ()  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1770  | 
(* if the model specifies a depth for the current type, *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1771  | 
(* decrement it to avoid infinite recursion *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1772  | 
val typs' = case depth of NONE => typs | SOME n =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1773  | 
AList.update (op =) (Type (s, Ts), n-1) typs  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1774  | 
(* recursively compute the size of the datatype *)  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1775  | 
val size = size_of_dtyp ctxt typs' descr typ_assoc constrs  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1776  | 
val next_idx = #next_idx args  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1777  | 
val next = next_idx+size  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1778  | 
(* check if 'maxvars' is large enough *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1779  | 
val _ = (if next-1 > #maxvars args andalso  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1780  | 
#maxvars args > 0 then raise MAXVARS_EXCEEDED else ())  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1781  | 
val fms = map BoolVar (next_idx upto (next_idx+size-1))  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1782  | 
val intr = Leaf fms  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1783  | 
fun one_of_two_false [] = True  | 
| 41471 | 1784  | 
| one_of_two_false (x::xs) = SAnd (Prop_Logic.all (map (fn x' =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1785  | 
SOr (SNot x, SNot x')) xs), one_of_two_false xs)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1786  | 
val wf = one_of_two_false fms  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1787  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1788  | 
(* extend the model, increase 'next_idx', add well-formedness *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1789  | 
(* condition *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1790  | 
                    SOME (intr, (typs, (t, intr)::terms), {maxvars = #maxvars args,
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1791  | 
def_eq = #def_eq args, next_idx = next, bounds = #bounds args,  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1792  | 
wellformed = SAnd (#wellformed args, wf)})  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1793  | 
end  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1794  | 
end  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1795  | 
| NONE => (* not an inductive datatype *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1796  | 
NONE)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1797  | 
| interpret_term _ = (* a (free or schematic) type variable *)  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1798  | 
NONE  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1799  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1800  | 
case AList.lookup (op =) terms t of  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1801  | 
SOME intr =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1802  | 
(* return an existing interpretation *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1803  | 
SOME (intr, model, args)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1804  | 
| NONE =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1805  | 
(case t of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1806  | 
Free (_, T) => interpret_term T  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1807  | 
| Var (_, T) => interpret_term T  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1808  | 
| Const (_, T) => interpret_term T  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1809  | 
| _ => NONE)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1810  | 
end;  | 
| 
15547
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
1811  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1812  | 
(* This function imposes an order on the elements of a datatype fragment *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1813  | 
(* as follows: C_i x_1 ... x_n < C_j y_1 ... y_m iff i < j or *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1814  | 
(* (x_1, ..., x_n) < (y_1, ..., y_m). With this order, a constructor is *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1815  | 
(* a function C_i that maps some argument indices x_1, ..., x_n to the *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1816  | 
(* datatype element given by index C_i x_1 ... x_n. The idea remains the *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1817  | 
(* same for recursive datatypes, although the computation of indices gets *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1818  | 
(* a little tricky. *)  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1819  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1820  | 
fun IDT_constructor_interpreter ctxt model args t =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1821  | 
let  | 
| 42361 | 1822  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1823  | 
(* returns a list of canonical representations for terms of the type 'T' *)  | 
| 
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1824  | 
(* It would be nice if we could just use 'print' for this, but 'print' *)  | 
| 
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1825  | 
(* for IDTs calls 'IDT_constructor_interpreter' again, and this could *)  | 
| 
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1826  | 
(* lead to infinite recursion when we have (mutually) recursive IDTs. *)  | 
| 
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1827  | 
fun canonical_terms typs T =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1828  | 
(case T of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1829  | 
            Type ("fun", [T1, T2]) =>
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1830  | 
(* 'T2' might contain a recursive IDT, so we cannot use 'print' (at *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1831  | 
(* least not for 'T2' *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1832  | 
let  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1833  | 
(* returns a list of lists, each one consisting of n (possibly *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1834  | 
(* identical) elements from 'xs' *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1835  | 
fun pick_all 1 xs = map single xs  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1836  | 
| pick_all n xs =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1837  | 
let val rec_pick = pick_all (n-1) xs in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1838  | 
maps (fn x => map (cons x) rec_pick) xs  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1839  | 
end  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1840  | 
(* ["x1", ..., "xn"] *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1841  | 
val terms1 = canonical_terms typs T1  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1842  | 
(* ["y1", ..., "ym"] *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1843  | 
val terms2 = canonical_terms typs T2  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1844  | 
              (* [[("x1", "y1"), ..., ("xn", "y1")], ..., *)
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1845  | 
              (*   [("x1", "ym"), ..., ("xn", "ym")]]     *)
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1846  | 
val functions = map (curry (op ~~) terms1)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1847  | 
(pick_all (length terms1) terms2)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1848  | 
(* [["(x1, y1)", ..., "(xn, y1)"], ..., *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1849  | 
(* ["(x1, ym)", ..., "(xn, ym)"]] *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1850  | 
val pairss = map (map HOLogic.mk_prod) functions  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1851  | 
val HOLogic_prodT = HOLogic.mk_prodT (T1, T2)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1852  | 
val HOLogic_setT = HOLogic.mk_setT HOLogic_prodT  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1853  | 
              val HOLogic_empty_set = Const (@{const_abbrev Set.empty}, HOLogic_setT)
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1854  | 
val HOLogic_insert =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1855  | 
                Const (@{const_name insert}, HOLogic_prodT --> HOLogic_setT --> HOLogic_setT)
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1856  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1857  | 
(* functions as graphs, i.e. as a (HOL) set of pairs "(x, y)" *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1858  | 
map (fn ps => fold_rev (fn pair => fn acc => HOLogic_insert $ pair $ acc) ps  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1859  | 
HOLogic_empty_set) pairss  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1860  | 
end  | 
| 
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1861  | 
| Type (s, Ts) =>  | 
| 
58354
 
04ac60da613e
support (finite values of) codatatypes in Quickcheck
 
blanchet 
parents: 
58322 
diff
changeset
 | 
1862  | 
(case BNF_LFP_Compat.get_info thy [] s of  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1863  | 
SOME info =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1864  | 
(case AList.lookup (op =) typs T of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1865  | 
SOME 0 =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1866  | 
(* termination condition to avoid infinite recursion *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1867  | 
[] (* at depth 0, every IDT is empty *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1868  | 
| _ =>  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1869  | 
let  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1870  | 
val index = #index info  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1871  | 
val descr = #descr info  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1872  | 
val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1873  | 
val typ_assoc = dtyps ~~ Ts  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1874  | 
(* sanity check: every element in 'dtyps' must be a 'DtTFree' *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1875  | 
val _ =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1876  | 
if Library.exists (fn d =>  | 
| 58120 | 1877  | 
case d of Old_Datatype_Aux.DtTFree _ => false | _ => true) dtyps  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1878  | 
then  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1879  | 
                      raise REFUTE ("IDT_constructor_interpreter",
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1880  | 
"datatype argument (for type "  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1881  | 
^ Syntax.string_of_typ ctxt T  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1882  | 
^ ") is not a variable")  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1883  | 
else ()  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1884  | 
(* decrement depth for the IDT 'T' *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1885  | 
val typs' =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1886  | 
(case AList.lookup (op =) typs T of NONE => typs  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1887  | 
| SOME n => AList.update (op =) (T, n-1) typs)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1888  | 
fun constructor_terms terms [] = terms  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1889  | 
| constructor_terms terms (d::ds) =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1890  | 
let  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1891  | 
val dT = typ_of_dtyp descr typ_assoc d  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1892  | 
val d_terms = canonical_terms typs' dT  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1893  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1894  | 
(* C_i x_1 ... x_n < C_i y_1 ... y_n if *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1895  | 
(* (x_1, ..., x_n) < (y_1, ..., y_n) *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1896  | 
constructor_terms  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1897  | 
(map_product (curry op $) terms d_terms) ds  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1898  | 
end  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1899  | 
in  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1900  | 
(* C_i ... < C_j ... if i < j *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1901  | 
maps (fn (cname, ctyps) =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1902  | 
let  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1903  | 
val cTerm = Const (cname,  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1904  | 
map (typ_of_dtyp descr typ_assoc) ctyps ---> T)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1905  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1906  | 
constructor_terms [cTerm] ctyps  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1907  | 
end) constrs  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1908  | 
end)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1909  | 
| NONE =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1910  | 
(* not an inductive datatype; in this case the argument types in *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1911  | 
(* 'Ts' may not be IDTs either, so 'print' should be safe *)  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1912  | 
map (fn intr => print ctxt (typs, []) T intr (K false))  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1913  | 
(make_constants ctxt (typs, []) T))  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1914  | 
| _ => (* TFree ..., TVar ... *)  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1915  | 
map (fn intr => print ctxt (typs, []) T intr (K false))  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1916  | 
(make_constants ctxt (typs, []) T))  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1917  | 
val (typs, terms) = model  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1918  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1919  | 
case AList.lookup (op =) terms t of  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1920  | 
SOME intr =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1921  | 
(* return an existing interpretation *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1922  | 
SOME (intr, model, args)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1923  | 
| NONE =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1924  | 
(case t of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1925  | 
Const (s, T) =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1926  | 
(case body_type T of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1927  | 
Type (s', Ts') =>  | 
| 
58354
 
04ac60da613e
support (finite values of) codatatypes in Quickcheck
 
blanchet 
parents: 
58322 
diff
changeset
 | 
1928  | 
(case BNF_LFP_Compat.get_info thy [] s' of  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1929  | 
SOME info => (* body type is an inductive datatype *)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1930  | 
let  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1931  | 
val index = #index info  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1932  | 
val descr = #descr info  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1933  | 
val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1934  | 
val typ_assoc = dtyps ~~ Ts'  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1935  | 
(* sanity check: every element in 'dtyps' must be a 'DtTFree' *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1936  | 
val _ = if Library.exists (fn d =>  | 
| 58120 | 1937  | 
case d of Old_Datatype_Aux.DtTFree _ => false | _ => true) dtyps  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1938  | 
then  | 
| 
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
1939  | 
                          raise REFUTE ("IDT_constructor_interpreter",
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1940  | 
"datatype argument (for type "  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1941  | 
^ Syntax.string_of_typ ctxt (Type (s', Ts'))  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1942  | 
^ ") is not a variable")  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
1943  | 
else ()  | 
| 58322 | 1944  | 
(* split the constructors into those occurring before/after *)  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1945  | 
(* 'Const (s, T)' *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1946  | 
val (constrs1, constrs2) = take_prefix (fn (cname, ctypes) =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1947  | 
not (cname = s andalso Sign.typ_instance thy (T,  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1948  | 
map (typ_of_dtyp descr typ_assoc) ctypes  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1949  | 
---> Type (s', Ts')))) constrs  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1950  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1951  | 
case constrs2 of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1952  | 
[] =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1953  | 
(* 'Const (s, T)' is not a constructor of this datatype *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1954  | 
NONE  | 
| 46096 | 1955  | 
| (_, ctypes)::_ =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1956  | 
let  | 
| 
55507
 
5f27fb2110e0
removed odd comments -- inferred types are shown by Prover IDE;
 
wenzelm 
parents: 
55411 
diff
changeset
 | 
1957  | 
(* only /recursive/ IDTs have an associated depth *)  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1958  | 
val depth = AList.lookup (op =) typs (Type (s', Ts'))  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1959  | 
(* this should never happen: at depth 0, this IDT fragment *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1960  | 
(* is definitely empty, and in this case we don't need to *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1961  | 
(* interpret its constructors *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1962  | 
val _ = (case depth of SOME 0 =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1963  | 
                                raise REFUTE ("IDT_constructor_interpreter",
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1964  | 
"depth is 0")  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1965  | 
| _ => ())  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1966  | 
val typs' = (case depth of NONE => typs | SOME n =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1967  | 
AList.update (op =) (Type (s', Ts'), n-1) typs)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1968  | 
(* elements of the datatype come before elements generated *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1969  | 
(* by 'Const (s, T)' iff they are generated by a *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1970  | 
(* constructor in constrs1 *)  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1971  | 
val offset = size_of_dtyp ctxt typs' descr typ_assoc constrs1  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1972  | 
(* compute the total (current) size of the datatype *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1973  | 
val total = offset +  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1974  | 
size_of_dtyp ctxt typs' descr typ_assoc constrs2  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1975  | 
(* sanity check *)  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1976  | 
val _ = if total <> size_of_type ctxt (typs, [])  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1977  | 
(Type (s', Ts')) then  | 
| 33246 | 1978  | 
                                raise REFUTE ("IDT_constructor_interpreter",
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1979  | 
"total is not equal to current size")  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1980  | 
else ()  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1981  | 
(* returns an interpretation where everything is mapped to *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1982  | 
(* an "undefined" element of the datatype *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1983  | 
fun make_undef [] = Leaf (replicate total False)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1984  | 
| make_undef (d::ds) =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1985  | 
let  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1986  | 
(* compute the current size of the type 'd' *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1987  | 
val dT = typ_of_dtyp descr typ_assoc d  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
1988  | 
val size = size_of_type ctxt (typs, []) dT  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1989  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1990  | 
Node (replicate size (make_undef ds))  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1991  | 
end  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1992  | 
(* returns the interpretation for a constructor *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1993  | 
fun make_constr [] offset =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1994  | 
if offset < total then  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1995  | 
(Leaf (replicate offset False @ True ::  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1996  | 
(replicate (total - offset - 1) False)), offset + 1)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1997  | 
else  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1998  | 
                                    raise REFUTE ("IDT_constructor_interpreter",
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
1999  | 
"offset >= total")  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2000  | 
| make_constr (d::ds) offset =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2001  | 
let  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2002  | 
val dT = typ_of_dtyp descr typ_assoc d  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2003  | 
(* compute canonical term representations for all *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2004  | 
(* elements of the type 'd' (with the reduced depth *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2005  | 
(* for the IDT) *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2006  | 
val terms' = canonical_terms typs' dT  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2007  | 
(* sanity check *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2008  | 
val _ =  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2009  | 
if length terms' <> size_of_type ctxt (typs', []) dT  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2010  | 
then  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2011  | 
                                        raise REFUTE ("IDT_constructor_interpreter",
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2012  | 
"length of terms' is not equal to old size")  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2013  | 
else ()  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2014  | 
(* compute canonical term representations for all *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2015  | 
(* elements of the type 'd' (with the current depth *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2016  | 
(* for the IDT) *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2017  | 
val terms = canonical_terms typs dT  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2018  | 
(* sanity check *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2019  | 
val _ =  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2020  | 
if length terms <> size_of_type ctxt (typs, []) dT  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2021  | 
then  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2022  | 
                                        raise REFUTE ("IDT_constructor_interpreter",
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2023  | 
"length of terms is not equal to current size")  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2024  | 
else ()  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2025  | 
(* sanity check *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2026  | 
val _ =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2027  | 
if length terms < length terms' then  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2028  | 
                                        raise REFUTE ("IDT_constructor_interpreter",
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2029  | 
"current size is less than old size")  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2030  | 
else ()  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2031  | 
(* sanity check: every element of terms' must also be *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2032  | 
(* present in terms *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2033  | 
val _ =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2034  | 
if forall (member (op =) terms) terms' then ()  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2035  | 
else  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2036  | 
                                        raise REFUTE ("IDT_constructor_interpreter",
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2037  | 
"element has disappeared")  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2038  | 
(* sanity check: the order on elements of terms' is *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2039  | 
(* the same in terms, for those elements *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2040  | 
val _ =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2041  | 
let  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2042  | 
fun search (x::xs) (y::ys) =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2043  | 
if x = y then search xs ys else search (x::xs) ys  | 
| 46096 | 2044  | 
| search (_::_) [] =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2045  | 
                                              raise REFUTE ("IDT_constructor_interpreter",
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2046  | 
"element order not preserved")  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2047  | 
| search [] _ = ()  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2048  | 
in search terms' terms end  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2049  | 
val (intrs, new_offset) =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2050  | 
fold_map (fn t_elem => fn off =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2051  | 
(* if 't_elem' existed at the previous depth, *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2052  | 
(* proceed recursively, otherwise map the entire *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2053  | 
(* subtree to "undefined" *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2054  | 
if member (op =) terms' t_elem then  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2055  | 
make_constr ds off  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2056  | 
else  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2057  | 
(make_undef ds, off))  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2058  | 
terms offset  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2059  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2060  | 
(Node intrs, new_offset)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2061  | 
end  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2062  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2063  | 
SOME (fst (make_constr ctypes offset), model, args)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2064  | 
end  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2065  | 
end  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2066  | 
| NONE => (* body type is not an inductive datatype *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2067  | 
NONE)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2068  | 
| _ => (* body type is a (free or schematic) type variable *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2069  | 
NONE)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2070  | 
| _ => (* term is not a constant *)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2071  | 
NONE)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2072  | 
end;  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
2073  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2074  | 
(* Difficult code ahead. Make sure you understand the *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2075  | 
(* 'IDT_constructor_interpreter' and the order in which it enumerates *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2076  | 
(* elements of an IDT before you try to understand this function. *)  | 
| 
15547
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
2077  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2078  | 
fun IDT_recursion_interpreter ctxt model args t =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2079  | 
let  | 
| 42361 | 2080  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2081  | 
in  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2082  | 
(* careful: here we descend arbitrarily deep into 't', possibly before *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2083  | 
(* any other interpreter for atomic terms has had a chance to look at *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2084  | 
(* 't' *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2085  | 
case strip_comb t of  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2086  | 
(Const (s, T), params) =>  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2087  | 
(* iterate over all datatypes in 'thy' *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2088  | 
Symtab.fold (fn (_, info) => fn result =>  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2089  | 
case result of  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2090  | 
SOME _ =>  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2091  | 
result (* just keep 'result' *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2092  | 
| NONE =>  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2093  | 
if member (op =) (#rec_names info) s then  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2094  | 
(* we do have a recursion operator of one of the (mutually *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2095  | 
(* recursive) datatypes given by 'info' *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2096  | 
let  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2097  | 
(* number of all constructors, including those of different *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2098  | 
(* (mutually recursive) datatypes within the same descriptor *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2099  | 
val mconstrs_count =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2100  | 
Integer.sum (map (fn (_, (_, _, cs)) => length cs) (#descr info))  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2101  | 
in  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2102  | 
if mconstrs_count < length params then  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2103  | 
(* too many actual parameters; for now we'll use the *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2104  | 
(* 'stlc_interpreter' to strip off one application *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2105  | 
NONE  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2106  | 
else if mconstrs_count > length params then  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2107  | 
(* too few actual parameters; we use eta expansion *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2108  | 
(* Note that the resulting expansion of lambda abstractions *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2109  | 
(* by the 'stlc_interpreter' may be rather slow (depending *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2110  | 
(* on the argument types and the size of the IDT, of *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2111  | 
(* course). *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2112  | 
SOME (interpret ctxt model args (eta_expand t  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2113  | 
(mconstrs_count - length params)))  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2114  | 
else (* mconstrs_count = length params *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2115  | 
let  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2116  | 
(* interpret each parameter separately *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2117  | 
val (p_intrs, (model', args')) = fold_map (fn p => fn (m, a) =>  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2118  | 
let  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2119  | 
val (i, m', a') = interpret ctxt m a p  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2120  | 
in  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2121  | 
(i, (m', a'))  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2122  | 
end) params (model, args)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2123  | 
val (typs, _) = model'  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2124  | 
(* 'index' is /not/ necessarily the index of the IDT that *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2125  | 
(* the recursion operator is associated with, but merely *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2126  | 
(* the index of some mutually recursive IDT *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2127  | 
val index = #index info  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2128  | 
val descr = #descr info  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2129  | 
val (_, dtyps, _) = the (AList.lookup (op =) descr index)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2130  | 
(* sanity check: we assume that the order of constructors *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2131  | 
(* in 'descr' is the same as the order of *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2132  | 
(* corresponding parameters, otherwise the *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2133  | 
(* association code below won't match the *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2134  | 
(* right constructors/parameters; we also *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2135  | 
(* assume that the order of recursion *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2136  | 
(* operators in '#rec_names info' is the *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2137  | 
(* same as the order of corresponding *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2138  | 
(* datatypes in 'descr' *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2139  | 
val _ = if map fst descr <> (0 upto (length descr - 1)) then  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2140  | 
                          raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2141  | 
"order of constructors and corresponding parameters/" ^  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2142  | 
"recursion operators and corresponding datatypes " ^  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2143  | 
"different?")  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2144  | 
else ()  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2145  | 
(* sanity check: every element in 'dtyps' must be a *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2146  | 
(* 'DtTFree' *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2147  | 
val _ =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2148  | 
if Library.exists (fn d =>  | 
| 58120 | 2149  | 
case d of Old_Datatype_Aux.DtTFree _ => false  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2150  | 
| _ => true) dtyps  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2151  | 
then  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2152  | 
                          raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2153  | 
"datatype argument is not a variable")  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2154  | 
else ()  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2155  | 
(* the type of a recursion operator is *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2156  | 
(* [T1, ..., Tn, IDT] ---> Tresult *)  | 
| 42364 | 2157  | 
val IDT = nth (binder_types T) mconstrs_count  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2158  | 
(* by our assumption on the order of recursion operators *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2159  | 
(* and datatypes, this is the index of the datatype *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2160  | 
(* corresponding to the given recursion operator *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2161  | 
val idt_index = find_index (fn s' => s' = s) (#rec_names info)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2162  | 
(* mutually recursive types must have the same type *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2163  | 
(* parameters, unless the mutual recursion comes from *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2164  | 
(* indirect recursion *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2165  | 
fun rec_typ_assoc acc [] = acc  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2166  | 
| rec_typ_assoc acc ((d, T)::xs) =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2167  | 
(case AList.lookup op= acc d of  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2168  | 
NONE =>  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2169  | 
(case d of  | 
| 58120 | 2170  | 
Old_Datatype_Aux.DtTFree _ =>  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2171  | 
(* add the association, proceed *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2172  | 
rec_typ_assoc ((d, T)::acc) xs  | 
| 58120 | 2173  | 
| Old_Datatype_Aux.DtType (s, ds) =>  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2174  | 
let  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2175  | 
val (s', Ts) = dest_Type T  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2176  | 
in  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2177  | 
if s=s' then  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2178  | 
rec_typ_assoc ((d, T)::acc) ((ds ~~ Ts) @ xs)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2179  | 
else  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2180  | 
                                        raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2181  | 
"DtType/Type mismatch")  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2182  | 
end  | 
| 58120 | 2183  | 
| Old_Datatype_Aux.DtRec i =>  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2184  | 
let  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2185  | 
val (_, ds, _) = the (AList.lookup (op =) descr i)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2186  | 
val (_, Ts) = dest_Type T  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2187  | 
in  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2188  | 
rec_typ_assoc ((d, T)::acc) ((ds ~~ Ts) @ xs)  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2189  | 
end)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2190  | 
| SOME T' =>  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2191  | 
if T=T' then  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2192  | 
(* ignore the association since it's already *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2193  | 
(* present, proceed *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2194  | 
rec_typ_assoc acc xs  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2195  | 
else  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2196  | 
                                  raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2197  | 
"different type associations for the same dtyp"))  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2198  | 
val typ_assoc = filter  | 
| 58120 | 2199  | 
(fn (Old_Datatype_Aux.DtTFree _, _) => true | (_, _) => false)  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2200  | 
(rec_typ_assoc []  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2201  | 
(#2 (the (AList.lookup (op =) descr idt_index)) ~~ (snd o dest_Type) IDT))  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2202  | 
(* sanity check: typ_assoc must associate types to the *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2203  | 
(* elements of 'dtyps' (and only to those) *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2204  | 
val _ =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2205  | 
if not (eq_set (op =) (dtyps, map fst typ_assoc))  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2206  | 
then  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2207  | 
                          raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2208  | 
"type association has extra/missing elements")  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2209  | 
else ()  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2210  | 
(* interpret each constructor in the descriptor (including *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2211  | 
(* those of mutually recursive datatypes) *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2212  | 
(* (int * interpretation list) list *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2213  | 
val mc_intrs = map (fn (idx, (_, _, cs)) =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2214  | 
let  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2215  | 
val c_return_typ = typ_of_dtyp descr typ_assoc  | 
| 58120 | 2216  | 
(Old_Datatype_Aux.DtRec idx)  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2217  | 
in  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2218  | 
(idx, map (fn (cname, cargs) =>  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2219  | 
                            (#1 o interpret ctxt (typs, []) {maxvars=0,
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2220  | 
def_eq=false, next_idx=1, bounds=[],  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2221  | 
wellformed=True}) (Const (cname, map (typ_of_dtyp  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2222  | 
descr typ_assoc) cargs ---> c_return_typ))) cs)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2223  | 
end) descr  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2224  | 
(* associate constructors with corresponding parameters *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2225  | 
(* (int * (interpretation * interpretation) list) list *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2226  | 
val (mc_p_intrs, p_intrs') = fold_map  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2227  | 
(fn (idx, c_intrs) => fn p_intrs' =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2228  | 
let  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2229  | 
val len = length c_intrs  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2230  | 
in  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2231  | 
((idx, c_intrs ~~ List.take (p_intrs', len)),  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2232  | 
List.drop (p_intrs', len))  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2233  | 
end) mc_intrs p_intrs  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2234  | 
(* sanity check: no 'p_intr' may be left afterwards *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2235  | 
val _ =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2236  | 
if p_intrs' <> [] then  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
2237  | 
                          raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2238  | 
"more parameter than constructor interpretations")  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2239  | 
else ()  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2240  | 
(* The recursion operator, applied to 'mconstrs_count' *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2241  | 
(* arguments, is a function that maps every element of the *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2242  | 
(* inductive datatype to an element of some result type. *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2243  | 
(* Recursion operators for mutually recursive IDTs are *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2244  | 
(* translated simultaneously. *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2245  | 
(* Since the order on datatype elements is given by an *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2246  | 
(* order on constructors (and then by the order on *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2247  | 
(* argument tuples), we can simply copy corresponding *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2248  | 
(* subtrees from 'p_intrs', in the order in which they are *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2249  | 
(* given. *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2250  | 
fun ci_pi (Leaf xs, pi) =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2251  | 
(* if the constructor does not match the arguments to a *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2252  | 
(* defined element of the IDT, the corresponding value *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2253  | 
(* of the parameter must be ignored *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2254  | 
if List.exists (equal True) xs then [pi] else []  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2255  | 
| ci_pi (Node xs, Node ys) = maps ci_pi (xs ~~ ys)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2256  | 
| ci_pi (Node _, Leaf _) =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2257  | 
                            raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2258  | 
"constructor takes more arguments than the " ^  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2259  | 
"associated parameter")  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2260  | 
val rec_operators = map (fn (idx, c_p_intrs) =>  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2261  | 
(idx, maps ci_pi c_p_intrs)) mc_p_intrs  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2262  | 
(* sanity check: every recursion operator must provide as *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2263  | 
(* many values as the corresponding datatype *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2264  | 
(* has elements *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2265  | 
val _ = map (fn (idx, intrs) =>  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2266  | 
let  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2267  | 
val T = typ_of_dtyp descr typ_assoc  | 
| 58120 | 2268  | 
(Old_Datatype_Aux.DtRec idx)  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2269  | 
in  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2270  | 
if length intrs <> size_of_type ctxt (typs, []) T then  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2271  | 
                            raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2272  | 
"wrong number of interpretations for rec. operator")  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2273  | 
else ()  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2274  | 
end) rec_operators  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2275  | 
(* For non-recursive datatypes, we are pretty much done at *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2276  | 
(* this point. For recursive datatypes however, we still *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2277  | 
(* need to apply the interpretations in 'rec_operators' to *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2278  | 
(* (recursively obtained) interpretations for recursive *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2279  | 
(* constructor arguments. To do so more efficiently, we *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2280  | 
(* copy 'rec_operators' into arrays first. Each Boolean *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2281  | 
(* indicates whether the recursive arguments have been *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2282  | 
(* considered already. *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2283  | 
val REC_OPERATORS = map (fn (idx, intrs) =>  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2284  | 
(idx, Array.fromList (map (pair false) intrs)))  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2285  | 
rec_operators  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2286  | 
(* takes an interpretation, and if some leaf of this *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2287  | 
(* interpretation is the 'elem'-th element of the type, *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2288  | 
(* the indices of the arguments leading to this leaf are *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2289  | 
(* returned *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2290  | 
fun get_args (Leaf xs) elem =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2291  | 
if find_index (fn x => x = True) xs = elem then  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2292  | 
SOME []  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2293  | 
else  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2294  | 
NONE  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2295  | 
| get_args (Node xs) elem =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2296  | 
let  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2297  | 
fun search ([], _) =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2298  | 
NONE  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2299  | 
| search (x::xs, n) =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2300  | 
(case get_args x elem of  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2301  | 
SOME result => SOME (n::result)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2302  | 
| NONE => search (xs, n+1))  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2303  | 
in  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2304  | 
search (xs, 0)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2305  | 
end  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2306  | 
(* returns the index of the constructor and indices for *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2307  | 
(* its arguments that generate the 'elem'-th element of *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2308  | 
(* the datatype given by 'idx' *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2309  | 
fun get_cargs idx elem =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2310  | 
let  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2311  | 
fun get_cargs_rec (_, []) =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2312  | 
                                raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2313  | 
"no matching constructor found for datatype element")  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2314  | 
| get_cargs_rec (n, x::xs) =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2315  | 
(case get_args x elem of  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2316  | 
SOME args => (n, args)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2317  | 
| NONE => get_cargs_rec (n+1, xs))  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2318  | 
in  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2319  | 
get_cargs_rec (0, the (AList.lookup (op =) mc_intrs idx))  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2320  | 
end  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2321  | 
(* computes one entry in 'REC_OPERATORS', and recursively *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2322  | 
(* all entries needed for it, where 'idx' gives the *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2323  | 
(* datatype and 'elem' the element of it *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2324  | 
fun compute_array_entry idx elem =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2325  | 
let  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2326  | 
val arr = the (AList.lookup (op =) REC_OPERATORS idx)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2327  | 
val (flag, intr) = Array.sub (arr, elem)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2328  | 
in  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2329  | 
if flag then  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2330  | 
(* simply return the previously computed result *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2331  | 
intr  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2332  | 
else  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2333  | 
(* we have to apply 'intr' to interpretations for all *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2334  | 
(* recursive arguments *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2335  | 
let  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2336  | 
val (c, args) = get_cargs idx elem  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2337  | 
(* find the indices of the constructor's /recursive/ *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2338  | 
(* arguments *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2339  | 
val (_, _, constrs) = the (AList.lookup (op =) descr idx)  | 
| 42364 | 2340  | 
val (_, dtyps) = nth constrs c  | 
2341  | 
val rec_dtyps_args = filter  | 
|
| 58120 | 2342  | 
(Old_Datatype_Aux.is_rec_type o fst) (dtyps ~~ args)  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2343  | 
(* map those indices to interpretations *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2344  | 
val rec_dtyps_intrs = map (fn (dtyp, arg) =>  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2345  | 
let  | 
| 42364 | 2346  | 
val dT = typ_of_dtyp descr typ_assoc dtyp  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2347  | 
val consts = make_constants ctxt (typs, []) dT  | 
| 42364 | 2348  | 
val arg_i = nth consts arg  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2349  | 
in  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2350  | 
(dtyp, arg_i)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2351  | 
end) rec_dtyps_args  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2352  | 
(* takes the dtyp and interpretation of an element, *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2353  | 
(* and computes the interpretation for the *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2354  | 
(* corresponding recursive argument *)  | 
| 58120 | 2355  | 
fun rec_intr (Old_Datatype_Aux.DtRec i) (Leaf xs) =  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2356  | 
(* recursive argument is "rec_i params elem" *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2357  | 
compute_array_entry i (find_index (fn x => x = True) xs)  | 
| 58120 | 2358  | 
| rec_intr (Old_Datatype_Aux.DtRec _) (Node _) =  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2359  | 
                                    raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2360  | 
"interpretation for IDT is a node")  | 
| 58120 | 2361  | 
                                | rec_intr (Old_Datatype_Aux.DtType ("fun", [_, dt2])) (Node xs) =
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2362  | 
(* recursive argument is something like *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2363  | 
(* "\<lambda>x::dt1. rec_? params (elem x)" *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2364  | 
Node (map (rec_intr dt2) xs)  | 
| 58120 | 2365  | 
                                | rec_intr (Old_Datatype_Aux.DtType ("fun", [_, _])) (Leaf _) =
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2366  | 
                                    raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2367  | 
"interpretation for function dtyp is a leaf")  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2368  | 
| rec_intr _ _ =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2369  | 
(* admissibility ensures that every recursive type *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2370  | 
(* is of the form 'Dt_1 -> ... -> Dt_k -> *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2371  | 
(* (DtRec i)' *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2372  | 
                                    raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2373  | 
"non-recursive codomain in recursive dtyp")  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2374  | 
(* obtain interpretations for recursive arguments *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2375  | 
(* interpretation list *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2376  | 
val arg_intrs = map (uncurry rec_intr) rec_dtyps_intrs  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2377  | 
(* apply 'intr' to all recursive arguments *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2378  | 
val result = fold (fn arg_i => fn i =>  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2379  | 
interpretation_apply (i, arg_i)) arg_intrs intr  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2380  | 
(* update 'REC_OPERATORS' *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2381  | 
val _ = Array.update (arr, elem, (true, result))  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2382  | 
in  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2383  | 
result  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2384  | 
end  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2385  | 
end  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2386  | 
val idt_size = Array.length (the (AList.lookup (op =) REC_OPERATORS idt_index))  | 
| 
56846
 
9df717fef2bb
renamed 'xxx_size' to 'size_xxx' for old datatype package
 
blanchet 
parents: 
56845 
diff
changeset
 | 
2387  | 
(* sanity check: the size of 'IDT' should be 'size_idt' *)  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2388  | 
val _ =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2389  | 
if idt_size <> size_of_type ctxt (typs, []) IDT then  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2390  | 
                            raise REFUTE ("IDT_recursion_interpreter",
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2391  | 
"unexpected size of IDT (wrong type associated?)")  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2392  | 
else ()  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2393  | 
val rec_op = Node (map_range (compute_array_entry idt_index) idt_size)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2394  | 
in  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2395  | 
SOME (rec_op, model', args')  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2396  | 
end  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2397  | 
end  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2398  | 
else  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2399  | 
NONE (* not a recursion operator of this datatype *)  | 
| 
58354
 
04ac60da613e
support (finite values of) codatatypes in Quickcheck
 
blanchet 
parents: 
58322 
diff
changeset
 | 
2400  | 
) (BNF_LFP_Compat.get_all thy []) NONE  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2401  | 
| _ => (* head of term is not a constant *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2402  | 
NONE  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2403  | 
end;  | 
| 
15547
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
2404  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2405  | 
fun set_interpreter ctxt model args t =  | 
| 
29956
 
62f931b257b7
Reintroduce set_interpreter for Collect and op :.
 
blanchet 
parents: 
29820 
diff
changeset
 | 
2406  | 
let  | 
| 
46098
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2407  | 
val (typs, terms) = model  | 
| 
29956
 
62f931b257b7
Reintroduce set_interpreter for Collect and op :.
 
blanchet 
parents: 
29820 
diff
changeset
 | 
2408  | 
in  | 
| 
 
62f931b257b7
Reintroduce set_interpreter for Collect and op :.
 
blanchet 
parents: 
29820 
diff
changeset
 | 
2409  | 
case AList.lookup (op =) terms t of  | 
| 
 
62f931b257b7
Reintroduce set_interpreter for Collect and op :.
 
blanchet 
parents: 
29820 
diff
changeset
 | 
2410  | 
SOME intr =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2411  | 
(* return an existing interpretation *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2412  | 
SOME (intr, model, args)  | 
| 
29956
 
62f931b257b7
Reintroduce set_interpreter for Collect and op :.
 
blanchet 
parents: 
29820 
diff
changeset
 | 
2413  | 
| NONE =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2414  | 
(case t of  | 
| 
46098
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2415  | 
          Free (x, Type (@{type_name set}, [T])) =>
 | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2416  | 
let  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2417  | 
val (intr, _, args') =  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2418  | 
interpret ctxt (typs, []) args (Free (x, T --> HOLogic.boolT))  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2419  | 
in  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2420  | 
SOME (intr, (typs, (t, intr)::terms), args')  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2421  | 
end  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2422  | 
        | Var ((x, i), Type (@{type_name set}, [T])) =>
 | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2423  | 
let  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2424  | 
val (intr, _, args') =  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2425  | 
interpret ctxt (typs, []) args (Var ((x,i), T --> HOLogic.boolT))  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2426  | 
in  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2427  | 
SOME (intr, (typs, (t, intr)::terms), args')  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2428  | 
end  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2429  | 
        | Const (s, Type (@{type_name set}, [T])) =>
 | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2430  | 
let  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2431  | 
val (intr, _, args') =  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2432  | 
interpret ctxt (typs, []) args (Const (s, T --> HOLogic.boolT))  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2433  | 
in  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2434  | 
SOME (intr, (typs, (t, intr)::terms), args')  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2435  | 
end  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2436  | 
(* 'Collect' == identity *)  | 
| 
46098
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2437  | 
        | Const (@{const_name Collect}, _) $ t1 =>
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2438  | 
SOME (interpret ctxt model args t1)  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2439  | 
        | Const (@{const_name Collect}, _) =>
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2440  | 
SOME (interpret ctxt model args (eta_expand t 1))  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2441  | 
(* 'op :' == application *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2442  | 
        | Const (@{const_name Set.member}, _) $ t1 $ t2 =>
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2443  | 
SOME (interpret ctxt model args (t2 $ t1))  | 
| 46096 | 2444  | 
        | Const (@{const_name Set.member}, _) $ _ =>
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2445  | 
SOME (interpret ctxt model args (eta_expand t 1))  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2446  | 
        | Const (@{const_name Set.member}, _) =>
 | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2447  | 
SOME (interpret ctxt model args (eta_expand t 2))  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2448  | 
| _ => NONE)  | 
| 
29956
 
62f931b257b7
Reintroduce set_interpreter for Collect and op :.
 
blanchet 
parents: 
29820 
diff
changeset
 | 
2449  | 
end;  | 
| 
 
62f931b257b7
Reintroduce set_interpreter for Collect and op :.
 
blanchet 
parents: 
29820 
diff
changeset
 | 
2450  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2451  | 
(* only an optimization: 'card' could in principle be interpreted with *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2452  | 
(* interpreters available already (using its definition), but the code *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2453  | 
(* below is more efficient *)  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
2454  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2455  | 
fun Finite_Set_card_interpreter ctxt model args t =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2456  | 
case t of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2457  | 
    Const (@{const_name Finite_Set.card},
 | 
| 
46098
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2458  | 
        Type ("fun", [Type (@{type_name set}, [T]), @{typ nat}])) =>
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2459  | 
let  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2460  | 
fun number_of_elements (Node xs) =  | 
| 33246 | 2461  | 
fold (fn x => fn n =>  | 
2462  | 
if x = TT then  | 
|
2463  | 
n + 1  | 
|
2464  | 
else if x = FF then  | 
|
2465  | 
n  | 
|
2466  | 
else  | 
|
2467  | 
                raise REFUTE ("Finite_Set_card_interpreter",
 | 
|
2468  | 
"interpretation for set type does not yield a Boolean"))  | 
|
2469  | 
xs 0  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2470  | 
| number_of_elements (Leaf _) =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2471  | 
              raise REFUTE ("Finite_Set_card_interpreter",
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2472  | 
"interpretation for set type is a leaf")  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2473  | 
        val size_of_nat = size_of_type ctxt model (@{typ nat})
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2474  | 
(* takes an interpretation for a set and returns an interpretation *)  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
2475  | 
(* for a 'nat' denoting the set's cardinality *)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2476  | 
fun card i =  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2477  | 
let  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2478  | 
val n = number_of_elements i  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2479  | 
in  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2480  | 
if n < size_of_nat then  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2481  | 
Leaf ((replicate n False) @ True ::  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
2482  | 
(replicate (size_of_nat-n-1) False))  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2483  | 
else  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
2484  | 
Leaf (replicate size_of_nat False)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2485  | 
end  | 
| 
46098
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2486  | 
val set_constants = make_constants ctxt model (HOLogic.mk_setT T)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2487  | 
in  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
2488  | 
SOME (Node (map card set_constants), model, args)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2489  | 
end  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2490  | 
| _ => NONE;  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
2491  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2492  | 
(* only an optimization: 'finite' could in principle be interpreted with *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2493  | 
(* interpreters available already (using its definition), but the code *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2494  | 
(* below is more efficient *)  | 
| 22093 | 2495  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2496  | 
fun Finite_Set_finite_interpreter ctxt model args t =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2497  | 
case t of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2498  | 
    Const (@{const_name Finite_Set.finite},
 | 
| 
46098
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2499  | 
           Type ("fun", [_, @{typ bool}])) $ _ =>
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2500  | 
(* we only consider finite models anyway, hence EVERY set is *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2501  | 
(* "finite" *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2502  | 
SOME (TT, model, args)  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2503  | 
  | Const (@{const_name Finite_Set.finite},
 | 
| 
46098
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2504  | 
           Type ("fun", [set_T, @{typ bool}])) =>
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2505  | 
let  | 
| 
46098
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2506  | 
val size_of_set = size_of_type ctxt model set_T  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2507  | 
in  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2508  | 
(* we only consider finite models anyway, hence EVERY set is *)  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2509  | 
(* "finite" *)  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
2510  | 
SOME (Node (replicate size_of_set TT), model, args)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2511  | 
end  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2512  | 
| _ => NONE;  | 
| 22093 | 2513  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2514  | 
(* only an optimization: 'less' could in principle be interpreted with *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2515  | 
(* interpreters available already (using its definition), but the code *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2516  | 
(* below is more efficient *)  | 
| 
15547
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
2517  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2518  | 
fun Nat_less_interpreter ctxt model args t =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2519  | 
case t of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2520  | 
    Const (@{const_name Orderings.less}, Type ("fun", [@{typ nat},
 | 
| 
38553
 
56965d8e4e11
use HOLogic.boolT and @{typ bool} more pervasively
 
haftmann 
parents: 
37677 
diff
changeset
 | 
2521  | 
        Type ("fun", [@{typ nat}, @{typ bool}])])) =>
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2522  | 
let  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2523  | 
        val size_of_nat = size_of_type ctxt model (@{typ nat})
 | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
2524  | 
(* the 'n'-th nat is not less than the first 'n' nats, while it *)  | 
| 
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
2525  | 
(* is less than the remaining 'size_of_nat - n' nats *)  | 
| 
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
2526  | 
fun less n = Node ((replicate n FF) @ (replicate (size_of_nat - n) TT))  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2527  | 
in  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
2528  | 
SOME (Node (map less (1 upto size_of_nat)), model, args)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2529  | 
end  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2530  | 
| _ => NONE;  | 
| 
15547
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
2531  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2532  | 
(* only an optimization: 'plus' could in principle be interpreted with *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2533  | 
(* interpreters available already (using its definition), but the code *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2534  | 
(* below is more efficient *)  | 
| 
15547
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
2535  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2536  | 
fun Nat_plus_interpreter ctxt model args t =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2537  | 
case t of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2538  | 
    Const (@{const_name Groups.plus}, Type ("fun", [@{typ nat},
 | 
| 37388 | 2539  | 
        Type ("fun", [@{typ nat}, @{typ nat}])])) =>
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2540  | 
let  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2541  | 
        val size_of_nat = size_of_type ctxt model (@{typ nat})
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2542  | 
fun plus m n =  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2543  | 
let  | 
| 25032 | 2544  | 
val element = m + n  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2545  | 
in  | 
| 25032 | 2546  | 
if element > size_of_nat - 1 then  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
2547  | 
Leaf (replicate size_of_nat False)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2548  | 
else  | 
| 25032 | 2549  | 
Leaf ((replicate element False) @ True ::  | 
2550  | 
(replicate (size_of_nat - element - 1) False))  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2551  | 
end  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2552  | 
in  | 
| 33063 | 2553  | 
SOME (Node (map_range (fn m => Node (map_range (plus m) size_of_nat)) size_of_nat),  | 
2554  | 
model, args)  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2555  | 
end  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2556  | 
| _ => NONE;  | 
| 
15547
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
2557  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2558  | 
(* only an optimization: 'minus' could in principle be interpreted *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2559  | 
(* with interpreters available already (using its definition), but the *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2560  | 
(* code below is more efficient *)  | 
| 
15547
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
2561  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2562  | 
fun Nat_minus_interpreter ctxt model args t =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2563  | 
case t of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2564  | 
    Const (@{const_name Groups.minus}, Type ("fun", [@{typ nat},
 | 
| 37388 | 2565  | 
        Type ("fun", [@{typ nat}, @{typ nat}])])) =>
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2566  | 
let  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2567  | 
        val size_of_nat = size_of_type ctxt model (@{typ nat})
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2568  | 
fun minus m n =  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2569  | 
let  | 
| 25032 | 2570  | 
val element = Int.max (m-n, 0)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2571  | 
in  | 
| 25032 | 2572  | 
Leaf ((replicate element False) @ True ::  | 
2573  | 
(replicate (size_of_nat - element - 1) False))  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2574  | 
end  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2575  | 
in  | 
| 33063 | 2576  | 
SOME (Node (map_range (fn m => Node (map_range (minus m) size_of_nat)) size_of_nat),  | 
2577  | 
model, args)  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2578  | 
end  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2579  | 
| _ => NONE;  | 
| 
15547
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
2580  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2581  | 
(* only an optimization: 'times' could in principle be interpreted *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2582  | 
(* with interpreters available already (using its definition), but the *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2583  | 
(* code below is more efficient *)  | 
| 
15547
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
2584  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2585  | 
fun Nat_times_interpreter ctxt model args t =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2586  | 
case t of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2587  | 
    Const (@{const_name Groups.times}, Type ("fun", [@{typ nat},
 | 
| 37388 | 2588  | 
        Type ("fun", [@{typ nat}, @{typ nat}])])) =>
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2589  | 
let  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2590  | 
        val size_of_nat = size_of_type ctxt model (@{typ nat})
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2591  | 
fun mult m n =  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2592  | 
let  | 
| 25032 | 2593  | 
val element = m * n  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2594  | 
in  | 
| 25032 | 2595  | 
if element > size_of_nat - 1 then  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
2596  | 
Leaf (replicate size_of_nat False)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2597  | 
else  | 
| 25032 | 2598  | 
Leaf ((replicate element False) @ True ::  | 
2599  | 
(replicate (size_of_nat - element - 1) False))  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2600  | 
end  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2601  | 
in  | 
| 33063 | 2602  | 
SOME (Node (map_range (fn m => Node (map_range (mult m) size_of_nat)) size_of_nat),  | 
2603  | 
model, args)  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2604  | 
end  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2605  | 
| _ => NONE;  | 
| 
15547
 
f08e2d83681e
major code change: refute can now handle recursion and axiomatic type classes; 3-valued logic with two kinds of equality; some bugfixes
 
webertj 
parents: 
15531 
diff
changeset
 | 
2606  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2607  | 
(* only an optimization: 'append' could in principle be interpreted with *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2608  | 
(* interpreters available already (using its definition), but the code *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2609  | 
(* below is more efficient *)  | 
| 
15767
 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 
webertj 
parents: 
15611 
diff
changeset
 | 
2610  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2611  | 
fun List_append_interpreter ctxt model args t =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2612  | 
case t of  | 
| 56252 | 2613  | 
    Const (@{const_name append},
 | 
2614  | 
      Type (@{type_name fun}, [Type (@{type_name list}, [T]),
 | 
|
2615  | 
        Type (@{type_name fun}, [Type (@{type_name list}, [_]), Type (@{type_name list}, [_])])])) =>
 | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2616  | 
let  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2617  | 
val size_elem = size_of_type ctxt model T  | 
| 56252 | 2618  | 
        val size_list = size_of_type ctxt model (Type (@{type_name list}, [T]))
 | 
| 25032 | 2619  | 
(* maximal length of lists; 0 if we only consider the empty list *)  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2620  | 
val list_length =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2621  | 
let  | 
| 25032 | 2622  | 
fun list_length_acc len lists total =  | 
2623  | 
if lists = total then  | 
|
2624  | 
len  | 
|
2625  | 
else if lists < total then  | 
|
2626  | 
list_length_acc (len+1) (lists*size_elem) (total-lists)  | 
|
2627  | 
else  | 
|
2628  | 
                raise REFUTE ("List_append_interpreter",
 | 
|
2629  | 
"size_list not equal to 1 + size_elem + ... + " ^  | 
|
2630  | 
"size_elem^len, for some len")  | 
|
2631  | 
in  | 
|
2632  | 
list_length_acc 0 1 size_list  | 
|
2633  | 
end  | 
|
2634  | 
val elements = 0 upto (size_list-1)  | 
|
2635  | 
(* FIXME: there should be a nice formula, which computes the same as *)  | 
|
2636  | 
(* the following, but without all this intermediate tree *)  | 
|
2637  | 
(* length/offset stuff *)  | 
|
2638  | 
(* associate each list with its length and offset in a complete tree *)  | 
|
2639  | 
(* of width 'size_elem' and depth 'length_list' (with 'size_list' *)  | 
|
2640  | 
(* nodes total) *)  | 
|
2641  | 
(* (int * (int * int)) list *)  | 
|
| 33246 | 2642  | 
val (lenoff_lists, _) = fold_map (fn elem => fn (offsets, off) =>  | 
| 25032 | 2643  | 
(* corresponds to a pre-order traversal of the tree *)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2644  | 
let  | 
| 25032 | 2645  | 
val len = length offsets  | 
2646  | 
(* associate the given element with len/off *)  | 
|
2647  | 
val assoc = (elem, (len, off))  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2648  | 
in  | 
| 25032 | 2649  | 
if len < list_length then  | 
2650  | 
(* go to first child node *)  | 
|
| 33246 | 2651  | 
(assoc, (off :: offsets, off * size_elem))  | 
| 25032 | 2652  | 
else if off mod size_elem < size_elem - 1 then  | 
2653  | 
(* go to next sibling node *)  | 
|
| 33246 | 2654  | 
(assoc, (offsets, off + 1))  | 
| 25032 | 2655  | 
else  | 
2656  | 
(* go back up the stack until we find a level where we can go *)  | 
|
2657  | 
(* to the next sibling node *)  | 
|
2658  | 
let  | 
|
| 
48902
 
44a6967240b7
prefer classic take_prefix/take_suffix over chop_while (cf. 0659e84bdc5f);
 
wenzelm 
parents: 
46961 
diff
changeset
 | 
2659  | 
val offsets' = snd (take_prefix  | 
| 39811 | 2660  | 
(fn off' => off' mod size_elem = size_elem - 1) offsets)  | 
| 25032 | 2661  | 
in  | 
2662  | 
case offsets' of  | 
|
2663  | 
[] =>  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2664  | 
(* we're at the last node in the tree; the next value *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2665  | 
(* won't be used anyway *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2666  | 
(assoc, ([], 0))  | 
| 25032 | 2667  | 
| off'::offs' =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2668  | 
(* go to next sibling node *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2669  | 
(assoc, (offs', off' + 1))  | 
| 25032 | 2670  | 
end  | 
| 33246 | 2671  | 
end) elements ([], 0)  | 
| 25032 | 2672  | 
(* we also need the reverse association (from length/offset to *)  | 
2673  | 
(* index) *)  | 
|
2674  | 
val lenoff'_lists = map Library.swap lenoff_lists  | 
|
2675  | 
(* returns the interpretation for "(list no. m) @ (list no. n)" *)  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2676  | 
fun append m n =  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2677  | 
let  | 
| 29288 | 2678  | 
val (len_m, off_m) = the (AList.lookup (op =) lenoff_lists m)  | 
2679  | 
val (len_n, off_n) = the (AList.lookup (op =) lenoff_lists n)  | 
|
| 25032 | 2680  | 
val len_elem = len_m + len_n  | 
| 
39047
 
cdff476ba39e
use existing Integer.pow, despite its slightly odd argument order;
 
wenzelm 
parents: 
39046 
diff
changeset
 | 
2681  | 
val off_elem = off_m * Integer.pow len_n size_elem + off_n  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2682  | 
in  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2683  | 
case AList.lookup op= lenoff'_lists (len_elem, off_elem) of  | 
| 25032 | 2684  | 
NONE =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2685  | 
(* undefined *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2686  | 
Leaf (replicate size_list False)  | 
| 25032 | 2687  | 
| SOME element =>  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2688  | 
Leaf ((replicate element False) @ True ::  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2689  | 
(replicate (size_list - element - 1) False))  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2690  | 
end  | 
| 
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2691  | 
in  | 
| 25032 | 2692  | 
SOME (Node (map (fn m => Node (map (append m) elements)) elements),  | 
2693  | 
model, args)  | 
|
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2694  | 
end  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2695  | 
| _ => NONE;  | 
| 
15767
 
8ed9fcc004fe
support for recursion over mutually recursive IDTs
 
webertj 
parents: 
15611 
diff
changeset
 | 
2696  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2697  | 
(* only an optimization: 'fst' could in principle be interpreted with *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2698  | 
(* interpreters available already (using its definition), but the code *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2699  | 
(* below is more efficient *)  | 
| 21267 | 2700  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2701  | 
fun Product_Type_fst_interpreter ctxt model args t =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2702  | 
case t of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2703  | 
    Const (@{const_name fst}, Type ("fun", [Type (@{type_name Product_Type.prod}, [T, U]), _])) =>
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2704  | 
let  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2705  | 
val constants_T = make_constants ctxt model T  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2706  | 
val size_U = size_of_type ctxt model U  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2707  | 
in  | 
| 32952 | 2708  | 
SOME (Node (maps (replicate size_U) constants_T), model, args)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2709  | 
end  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2710  | 
| _ => NONE;  | 
| 21267 | 2711  | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2712  | 
(* only an optimization: 'snd' could in principle be interpreted with *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2713  | 
(* interpreters available already (using its definition), but the code *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2714  | 
(* below is more efficient *)  | 
| 21267 | 2715  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2716  | 
fun Product_Type_snd_interpreter ctxt model args t =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2717  | 
case t of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2718  | 
    Const (@{const_name snd}, Type ("fun", [Type (@{type_name Product_Type.prod}, [T, U]), _])) =>
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2719  | 
let  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2720  | 
val size_T = size_of_type ctxt model T  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2721  | 
val constants_U = make_constants ctxt model U  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2722  | 
in  | 
| 32952 | 2723  | 
SOME (Node (flat (replicate size_T constants_U)), model, args)  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2724  | 
end  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2725  | 
| _ => NONE;  | 
| 21267 | 2726  | 
|
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
2727  | 
|
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
2728  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
2729  | 
(* PRINTERS *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
2730  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
2731  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2732  | 
fun stlc_printer ctxt model T intr assignment =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2733  | 
let  | 
| 40720 | 2734  | 
val strip_leading_quote = perhaps (try (unprefix "'"))  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2735  | 
fun string_of_typ (Type (s, _)) = s  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2736  | 
| string_of_typ (TFree (x, _)) = strip_leading_quote x  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2737  | 
| string_of_typ (TVar ((x,i), _)) =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2738  | 
strip_leading_quote x ^ string_of_int i  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2739  | 
fun index_from_interpretation (Leaf xs) =  | 
| 41471 | 2740  | 
find_index (Prop_Logic.eval assignment) xs  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2741  | 
| index_from_interpretation _ =  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2742  | 
          raise REFUTE ("stlc_printer",
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2743  | 
"interpretation for ground type is not a leaf")  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2744  | 
in  | 
| 
25014
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
2745  | 
case T of  | 
| 
 
b711b0af178e
significant code overhaul, bugfix for inductive data types
 
webertj 
parents: 
24928 
diff
changeset
 | 
2746  | 
      Type ("fun", [T1, T2]) =>
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2747  | 
let  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2748  | 
(* create all constants of type 'T1' *)  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2749  | 
val constants = make_constants ctxt model T1  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2750  | 
val results =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2751  | 
(case intr of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2752  | 
Node xs => xs  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2753  | 
            | _ => raise REFUTE ("stlc_printer",
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2754  | 
"interpretation for function type is a leaf"))  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2755  | 
(* Term.term list *)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2756  | 
val pairs = map (fn (arg, result) =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2757  | 
HOLogic.mk_prod  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2758  | 
(print ctxt model T1 arg assignment,  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2759  | 
print ctxt model T2 result assignment))  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2760  | 
(constants ~~ results)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2761  | 
val HOLogic_prodT = HOLogic.mk_prodT (T1, T2)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2762  | 
val HOLogic_setT = HOLogic.mk_setT HOLogic_prodT  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2763  | 
          val HOLogic_empty_set = Const (@{const_abbrev Set.empty}, HOLogic_setT)
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2764  | 
val HOLogic_insert =  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2765  | 
            Const (@{const_name insert}, HOLogic_prodT --> HOLogic_setT --> HOLogic_setT)
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2766  | 
in  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2767  | 
SOME (fold_rev (fn pair => fn acc => HOLogic_insert $ pair $ acc) pairs HOLogic_empty_set)  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2768  | 
end  | 
| 56252 | 2769  | 
    | Type (@{type_name prop}, []) =>
 | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2770  | 
(case index_from_interpretation intr of  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2771  | 
          ~1 => SOME (HOLogic.mk_Trueprop (Const (@{const_name undefined}, HOLogic.boolT)))
 | 
| 45740 | 2772  | 
        | 0  => SOME (HOLogic.mk_Trueprop @{term True})
 | 
2773  | 
        | 1  => SOME (HOLogic.mk_Trueprop @{term False})
 | 
|
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2774  | 
        | _  => raise REFUTE ("stlc_interpreter",
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2775  | 
"illegal interpretation for a propositional value"))  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2776  | 
| Type _ =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2777  | 
if index_from_interpretation intr = (~1) then  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2778  | 
          SOME (Const (@{const_name undefined}, T))
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2779  | 
else  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2780  | 
SOME (Const (string_of_typ T ^  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2781  | 
string_of_int (index_from_interpretation intr), T))  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2782  | 
| TFree _ =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2783  | 
if index_from_interpretation intr = (~1) then  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2784  | 
          SOME (Const (@{const_name undefined}, T))
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2785  | 
else  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2786  | 
SOME (Const (string_of_typ T ^  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2787  | 
string_of_int (index_from_interpretation intr), T))  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2788  | 
| TVar _ =>  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2789  | 
if index_from_interpretation intr = (~1) then  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2790  | 
          SOME (Const (@{const_name undefined}, T))
 | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2791  | 
else  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2792  | 
SOME (Const (string_of_typ T ^  | 
| 
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2793  | 
string_of_int (index_from_interpretation intr), T))  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
22093 
diff
changeset
 | 
2794  | 
end;  | 
| 14350 | 2795  | 
|
| 
46098
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2796  | 
fun set_printer ctxt model T intr assignment =  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2797  | 
(case T of  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2798  | 
    Type (@{type_name set}, [T1]) =>
 | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2799  | 
let  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2800  | 
(* create all constants of type 'T1' *)  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2801  | 
val constants = make_constants ctxt model T1  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2802  | 
val results = (case intr of  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2803  | 
Node xs => xs  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2804  | 
        | _       => raise REFUTE ("set_printer",
 | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2805  | 
"interpretation for set type is a leaf"))  | 
| 57820 | 2806  | 
val elements = map_filter (fn (arg, result) =>  | 
| 
46098
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2807  | 
case result of  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2808  | 
Leaf [fmTrue, (* fmFalse *) _] =>  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2809  | 
if Prop_Logic.eval assignment fmTrue then  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2810  | 
SOME (print ctxt model T1 arg assignment)  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2811  | 
else (* if Prop_Logic.eval assignment fmFalse then *)  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2812  | 
NONE  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2813  | 
| _ =>  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2814  | 
          raise REFUTE ("set_printer",
 | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2815  | 
"illegal interpretation for a Boolean value"))  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2816  | 
(constants ~~ results)  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2817  | 
val HOLogic_setT1 = HOLogic.mk_setT T1  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2818  | 
      val HOLogic_empty_set = Const (@{const_abbrev Set.empty}, HOLogic_setT1)
 | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2819  | 
val HOLogic_insert =  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2820  | 
        Const (@{const_name insert}, T1 --> HOLogic_setT1 --> HOLogic_setT1)
 | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2821  | 
in  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2822  | 
SOME (Library.foldl (fn (acc, elem) => HOLogic_insert $ elem $ acc)  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2823  | 
(HOLogic_empty_set, elements))  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2824  | 
end  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2825  | 
| _ =>  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2826  | 
NONE);  | 
| 
 
ce939621a1fe
handle "set" correctly in Refute -- inspired by old code from Isabelle2007
 
blanchet 
parents: 
46096 
diff
changeset
 | 
2827  | 
|
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2828  | 
fun IDT_printer ctxt model T intr assignment =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2829  | 
let  | 
| 42361 | 2830  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2831  | 
in  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2832  | 
(case T of  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2833  | 
Type (s, Ts) =>  | 
| 
58354
 
04ac60da613e
support (finite values of) codatatypes in Quickcheck
 
blanchet 
parents: 
58322 
diff
changeset
 | 
2834  | 
(case BNF_LFP_Compat.get_info thy [] s of  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2835  | 
SOME info => (* inductive datatype *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2836  | 
let  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2837  | 
val (typs, _) = model  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2838  | 
val index = #index info  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2839  | 
val descr = #descr info  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2840  | 
val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2841  | 
val typ_assoc = dtyps ~~ Ts  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2842  | 
(* sanity check: every element in 'dtyps' must be a 'DtTFree' *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2843  | 
val _ =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2844  | 
if Library.exists (fn d =>  | 
| 58120 | 2845  | 
case d of Old_Datatype_Aux.DtTFree _ => false | _ => true) dtyps  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2846  | 
then  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2847  | 
                  raise REFUTE ("IDT_printer", "datatype argument (for type " ^
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2848  | 
Syntax.string_of_typ ctxt (Type (s, Ts)) ^ ") is not a variable")  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2849  | 
else ()  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2850  | 
(* the index of the element in the datatype *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2851  | 
val element =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2852  | 
(case intr of  | 
| 41471 | 2853  | 
Leaf xs => find_index (Prop_Logic.eval assignment) xs  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2854  | 
                | Node _  => raise REFUTE ("IDT_printer",
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2855  | 
"interpretation is not a leaf"))  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2856  | 
in  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2857  | 
if element < 0 then  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2858  | 
                SOME (Const (@{const_name undefined}, Type (s, Ts)))
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2859  | 
else  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2860  | 
let  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2861  | 
(* takes a datatype constructor, and if for some arguments this *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2862  | 
(* constructor generates the datatype's element that is given by *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2863  | 
(* 'element', returns the constructor (as a term) as well as the *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2864  | 
(* indices of the arguments *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2865  | 
fun get_constr_args (cname, cargs) =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2866  | 
let  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2867  | 
val cTerm = Const (cname,  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2868  | 
map (typ_of_dtyp descr typ_assoc) cargs ---> Type (s, Ts))  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2869  | 
                      val (iC, _, _) = interpret ctxt (typs, []) {maxvars=0,
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2870  | 
def_eq=false, next_idx=1, bounds=[], wellformed=True} cTerm  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2871  | 
fun get_args (Leaf xs) =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2872  | 
if find_index (fn x => x = True) xs = element then  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2873  | 
SOME []  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2874  | 
else  | 
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2875  | 
NONE  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2876  | 
| get_args (Node xs) =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2877  | 
let  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2878  | 
fun search ([], _) =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2879  | 
NONE  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2880  | 
| search (x::xs, n) =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2881  | 
(case get_args x of  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2882  | 
SOME result => SOME (n::result)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2883  | 
| NONE => search (xs, n+1))  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2884  | 
in  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2885  | 
search (xs, 0)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2886  | 
end  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2887  | 
in  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2888  | 
Option.map (fn args => (cTerm, cargs, args)) (get_args iC)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2889  | 
end  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2890  | 
val (cTerm, cargs, args) =  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2891  | 
(* we could speed things up by computing the correct *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2892  | 
(* constructor directly (rather than testing all *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2893  | 
(* constructors), based on the order in which constructors *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2894  | 
(* generate elements of datatypes; the current implementation *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2895  | 
(* of 'IDT_printer' however is independent of the internals *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2896  | 
(* of 'IDT_constructor_interpreter' *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2897  | 
(case get_first get_constr_args constrs of  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2898  | 
SOME x => x  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2899  | 
                    | NONE   => raise REFUTE ("IDT_printer",
 | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2900  | 
"no matching constructor found for element " ^  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2901  | 
string_of_int element))  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2902  | 
val argsTerms = map (fn (d, n) =>  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2903  | 
let  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2904  | 
val dT = typ_of_dtyp descr typ_assoc d  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2905  | 
(* we only need the n-th element of this list, so there *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2906  | 
(* might be a more efficient implementation that does not *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2907  | 
(* generate all constants *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2908  | 
val consts = make_constants ctxt (typs, []) dT  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2909  | 
in  | 
| 42364 | 2910  | 
print ctxt (typs, []) dT (nth consts n) assignment  | 
| 
39049
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2911  | 
end) (cargs ~~ args)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2912  | 
in  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2913  | 
SOME (list_comb (cTerm, argsTerms))  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2914  | 
end  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2915  | 
end  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2916  | 
| NONE => (* not an inductive datatype *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2917  | 
NONE)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2918  | 
| _ => (* a (free or schematic) type variable *)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2919  | 
NONE)  | 
| 
 
423b72f2d242
prefer regular Proof.context over background theory;
 
wenzelm 
parents: 
39048 
diff
changeset
 | 
2920  | 
end;  | 
| 14350 | 2921  | 
|
2922  | 
||
2923  | 
(* ------------------------------------------------------------------------- *)  | 
|
| 
14456
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
2924  | 
(* Note: the interpreters and printers are used in reverse order; however, *)  | 
| 
 
cca28ec5f9a6
support for non-recursive IDTs, The, arbitrary, Hilbert_Choice.Eps
 
webertj 
parents: 
14351 
diff
changeset
 | 
2925  | 
(* an interpreter that can handle non-atomic terms ends up being *)  | 
| 
14807
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
2926  | 
(* applied before the 'stlc_interpreter' breaks the term apart into *)  | 
| 
 
e8ccb13d7774
major code change: refute can now handle any Isabelle term, adds certain axioms automatically, and can handle inductive datatypes (but not yet recursion over them)
 
webertj 
parents: 
14604 
diff
changeset
 | 
2927  | 
(* subterms that are then passed to other interpreters! *)  | 
| 14350 | 2928  | 
(* ------------------------------------------------------------------------- *)  | 
| 56256 | 2929  | 
(* FIXME formal @{const_name} for some entries (!??) *)
 | 
| 58825 | 2930  | 
val _ =  | 
2931  | 
Theory.setup  | 
|
2932  | 
(add_interpreter "stlc" stlc_interpreter #>  | 
|
2933  | 
add_interpreter "Pure" Pure_interpreter #>  | 
|
2934  | 
add_interpreter "HOLogic" HOLogic_interpreter #>  | 
|
2935  | 
add_interpreter "set" set_interpreter #>  | 
|
2936  | 
add_interpreter "IDT" IDT_interpreter #>  | 
|
2937  | 
add_interpreter "IDT_constructor" IDT_constructor_interpreter #>  | 
|
2938  | 
add_interpreter "IDT_recursion" IDT_recursion_interpreter #>  | 
|
2939  | 
add_interpreter "Finite_Set.card" Finite_Set_card_interpreter #>  | 
|
2940  | 
add_interpreter "Finite_Set.finite" Finite_Set_finite_interpreter #>  | 
|
2941  | 
add_interpreter "Nat_Orderings.less" Nat_less_interpreter #>  | 
|
2942  | 
add_interpreter "Nat_HOL.plus" Nat_plus_interpreter #>  | 
|
2943  | 
add_interpreter "Nat_HOL.minus" Nat_minus_interpreter #>  | 
|
2944  | 
add_interpreter "Nat_HOL.times" Nat_times_interpreter #>  | 
|
2945  | 
add_interpreter "List.append" List_append_interpreter #>  | 
|
2946  | 
add_interpreter "Product_Type.prod.fst" Product_Type_fst_interpreter #>  | 
|
2947  | 
add_interpreter "Product_Type.prod.snd" Product_Type_snd_interpreter #>  | 
|
2948  | 
add_printer "stlc" stlc_printer #>  | 
|
2949  | 
add_printer "set" set_printer #>  | 
|
2950  | 
add_printer "IDT" IDT_printer);  | 
|
| 14350 | 2951  | 
|
| 39048 | 2952  | 
|
2953  | 
||
2954  | 
(** outer syntax commands 'refute' and 'refute_params' **)  | 
|
2955  | 
||
2956  | 
(* argument parsing *)  | 
|
2957  | 
||
2958  | 
(*optional list of arguments of the form [name1=value1, name2=value2, ...]*)  | 
|
2959  | 
||
| 46949 | 2960  | 
val scan_parm = Parse.name -- (Scan.optional (@{keyword "="} |-- Parse.name) "true")
 | 
2961  | 
val scan_parms = Scan.optional (@{keyword "["} |-- Parse.list scan_parm --| @{keyword "]"}) [];
 | 
|
| 39048 | 2962  | 
|
2963  | 
||
2964  | 
(* 'refute' command *)  | 
|
2965  | 
||
2966  | 
val _ =  | 
|
| 
59936
 
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
 
wenzelm 
parents: 
59582 
diff
changeset
 | 
2967  | 
  Outer_Syntax.command @{command_keyword refute}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46960 
diff
changeset
 | 
2968  | 
"try to find a model that refutes a given subgoal"  | 
| 39048 | 2969  | 
(scan_parms -- Scan.optional Parse.nat 1 >>  | 
2970  | 
(fn (parms, i) =>  | 
|
| 
60190
 
906de96ba68a
allow diagnostic proof commands with skip_proofs;
 
wenzelm 
parents: 
60094 
diff
changeset
 | 
2971  | 
Toplevel.keep_proof (fn state =>  | 
| 39048 | 2972  | 
let  | 
2973  | 
val ctxt = Toplevel.context_of state;  | 
|
2974  | 
            val {goal = st, ...} = Proof.raw_goal (Toplevel.proof_of state);
 | 
|
| 
45387
 
ccffb3f9f42b
return outcome code, so that it can be picked up by Mutabelle
 
blanchet 
parents: 
44121 
diff
changeset
 | 
2975  | 
in refute_goal ctxt parms st i; () end)));  | 
| 39048 | 2976  | 
|
2977  | 
||
2978  | 
(* 'refute_params' command *)  | 
|
2979  | 
||
2980  | 
val _ =  | 
|
| 
59936
 
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
 
wenzelm 
parents: 
59582 
diff
changeset
 | 
2981  | 
  Outer_Syntax.command @{command_keyword refute_params}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46960 
diff
changeset
 | 
2982  | 
"show/store default parameters for the 'refute' command"  | 
| 39048 | 2983  | 
(scan_parms >> (fn parms =>  | 
2984  | 
Toplevel.theory (fn thy =>  | 
|
2985  | 
let  | 
|
2986  | 
val thy' = fold set_default_param parms thy;  | 
|
2987  | 
val output =  | 
|
| 42361 | 2988  | 
(case get_default_params (Proof_Context.init_global thy') of  | 
| 39048 | 2989  | 
[] => "none"  | 
2990  | 
| new_defaults => cat_lines (map (fn (x, y) => x ^ "=" ^ y) new_defaults));  | 
|
2991  | 
          val _ = writeln ("Default parameters for 'refute':\n" ^ output);
 | 
|
2992  | 
in thy' end)));  | 
|
2993  | 
||
| 
39046
 
5b38730f3e12
tuned whitespace and indentation, emphasizing the logical structure of this long text;
 
wenzelm 
parents: 
38864 
diff
changeset
 | 
2994  | 
end;  |