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