author | haftmann |
Wed, 23 Sep 2009 14:00:43 +0200 | |
changeset 32701 | 5059a733a4b8 |
parent 32657 | 5f13912245ff |
parent 32683 | 7c1fe854ca6a |
child 32740 | 9dd0a2f83429 |
permissions | -rw-r--r-- |
31217 | 1 |
(* Author: Lukas Bulwahn, TU Muenchen |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
2 |
|
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
3 |
(Prototype of) A compiler from predicates specified by intro/elim rules |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
4 |
to equations. |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
5 |
*) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
6 |
|
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
7 |
signature PREDICATE_COMPILE = |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
8 |
sig |
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
9 |
type mode = int list option list * int list |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
10 |
(*val add_equations_of: bool -> string list -> theory -> theory *) |
31514 | 11 |
val register_predicate : (thm list * thm * int) -> theory -> theory |
31877
e3de75d3b898
exported is_registered; added debug messages; removed extended parameter compilation in predicate compiler
bulwahn
parents:
31876
diff
changeset
|
12 |
val is_registered : theory -> string -> bool |
32314
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
13 |
(* val fetch_pred_data : theory -> string -> (thm list * thm * int) *) |
31514 | 14 |
val predfun_intro_of: theory -> string -> mode -> thm |
15 |
val predfun_elim_of: theory -> string -> mode -> thm |
|
16 |
val strip_intro_concl: int -> term -> term * (term list * term list) |
|
17 |
val predfun_name_of: theory -> string -> mode -> string |
|
18 |
val all_preds_of : theory -> string list |
|
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
19 |
val modes_of: theory -> string -> mode list |
32306
19f55947d4d5
removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents:
32287
diff
changeset
|
20 |
val string_of_mode : mode -> string |
31514 | 21 |
val intros_of: theory -> string -> thm list |
22 |
val nparams_of: theory -> string -> int |
|
31876
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
23 |
val add_intro: thm -> theory -> theory |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
24 |
val set_elim: thm -> theory -> theory |
31124 | 25 |
val setup: theory -> theory |
26 |
val code_pred: string -> Proof.context -> Proof.state |
|
27 |
val code_pred_cmd: string -> Proof.context -> Proof.state |
|
31876
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
28 |
val print_stored_rules: theory -> unit |
32306
19f55947d4d5
removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents:
32287
diff
changeset
|
29 |
val print_all_modes: theory -> unit |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
30 |
val do_proofs: bool ref |
31876
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
31 |
val mk_casesrule : Proof.context -> int -> thm list -> term |
31217 | 32 |
val analyze_compr: theory -> term -> term |
33 |
val eval_ref: (unit -> term Predicate.pred) option ref |
|
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
34 |
val add_equations : string list -> theory -> theory |
31876
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
35 |
val code_pred_intros_attrib : attribute |
32306
19f55947d4d5
removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents:
32287
diff
changeset
|
36 |
(* used by Quickcheck_Generator *) |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
37 |
(*val funT_of : mode -> typ -> typ |
32306
19f55947d4d5
removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents:
32287
diff
changeset
|
38 |
val mk_if_pred : term -> term |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
39 |
val mk_Eval : term * term -> term*) |
32307
55166cd57a6d
exported functions for quickcheck generator; renamed type constructing functions to fit with the type name in the Predicate theory
bulwahn
parents:
32306
diff
changeset
|
40 |
val mk_tupleT : typ list -> typ |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
41 |
(* val mk_predT : typ -> typ *) |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
42 |
(* temporary for testing of the compilation *) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
43 |
datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term | |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
44 |
GeneratorPrem of term list * term | Generator of (string * typ); |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
45 |
val prepare_intrs: theory -> string list -> |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
46 |
(string * typ) list * int * string list * string list * (string * mode list) list * |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
47 |
(string * (term list * indprem list) list) list * (string * (int option list * int)) list |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
48 |
datatype compilation_funs = CompilationFuns of { |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
49 |
mk_predT : typ -> typ, |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
50 |
dest_predT : typ -> typ, |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
51 |
mk_bot : typ -> term, |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
52 |
mk_single : term -> term, |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
53 |
mk_bind : term * term -> term, |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
54 |
mk_sup : term * term -> term, |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
55 |
mk_if : term -> term, |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
56 |
mk_not : term -> term, |
32351 | 57 |
mk_map : typ -> typ -> term -> term -> term, |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
58 |
lift_pred : term -> term |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
59 |
}; |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
60 |
datatype tmode = Mode of mode * int list * tmode option list; |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
61 |
type moded_clause = term list * (indprem * tmode) list |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
62 |
type 'a pred_mode_table = (string * (mode * 'a) list) list |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
63 |
val infer_modes : bool -> theory -> (string * (int list option list * int list) list) list |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
64 |
-> (string * (int option list * int)) list -> string list |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
65 |
-> (string * (term list * indprem list) list) list |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
66 |
-> (moded_clause list) pred_mode_table |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
67 |
val infer_modes_with_generator : theory -> (string * (int list option list * int list) list) list |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
68 |
-> (string * (int option list * int)) list -> string list |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
69 |
-> (string * (term list * indprem list) list) list |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
70 |
-> (moded_clause list) pred_mode_table |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
71 |
(*val compile_preds : theory -> compilation_funs -> string list -> string list |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
72 |
-> (string * typ) list -> (moded_clause list) pred_mode_table -> term pred_mode_table |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
73 |
val rpred_create_definitions :(string * typ) list -> string * mode list |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
74 |
-> theory -> theory |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
75 |
val split_smode : int list -> term list -> (term list * term list) *) |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
76 |
val print_moded_clauses : |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
77 |
theory -> (moded_clause list) pred_mode_table -> unit |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
78 |
val print_compiled_terms : theory -> term pred_mode_table -> unit |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
79 |
(*val rpred_prove_preds : theory -> term pred_mode_table -> thm pred_mode_table*) |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
80 |
val rpred_compfuns : compilation_funs |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
81 |
val dest_funT : typ -> typ * typ |
32314
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
82 |
(* val depending_preds_of : theory -> thm list -> string list *) |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
83 |
val add_quickcheck_equations : string list -> theory -> theory |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
84 |
val add_sizelim_equations : string list -> theory -> theory |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
85 |
val is_inductive_predicate : theory -> string -> bool |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
86 |
val terms_vs : term list -> string list |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
87 |
val subsets : int -> int -> int list list |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
88 |
val check_mode_clause : bool -> theory -> string list -> |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
89 |
(string * mode list) list -> (string * mode list) list -> mode -> (term list * indprem list) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
90 |
-> (term list * (indprem * tmode) list) option |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
91 |
val string_of_moded_prem : theory -> (indprem * tmode) -> string |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
92 |
val all_modes_of : theory -> (string * mode list) list |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
93 |
val all_generator_modes_of : theory -> (string * mode list) list |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
94 |
val compile_clause : compilation_funs -> term option -> (term list -> term) -> |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
95 |
theory -> string list -> string list -> mode -> term -> moded_clause -> term |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
96 |
val preprocess_intro : theory -> thm -> thm |
32315
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
97 |
val is_constrt : theory -> term -> bool |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
98 |
val is_predT : typ -> bool |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
99 |
val guess_nparams : typ -> int |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
100 |
end; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
101 |
|
31124 | 102 |
structure Predicate_Compile : PREDICATE_COMPILE = |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
103 |
struct |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
104 |
|
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
105 |
(** auxiliary **) |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
106 |
|
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
107 |
(* debug stuff *) |
31320
72eeb1b4e006
provide local dummy version of makestring -- NB: makestring is fragile and not portable, it should not occur in repository sources;
wenzelm
parents:
31284
diff
changeset
|
108 |
|
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
109 |
fun tracing s = (if ! Toplevel.debug then Output.tracing s else ()); |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
110 |
|
32314
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
111 |
fun print_tac s = Seq.single; (* (if ! Toplevel.debug then Tactical.print_tac s else Seq.single); *) |
32315
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
112 |
fun debug_tac msg = Seq.single; (* (fn st => (Output.tracing msg; Seq.single st)); *) |
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
113 |
|
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
114 |
val do_proofs = ref true; |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
115 |
|
31514 | 116 |
fun mycheat_tac thy i st = |
117 |
(Tactic.rtac (SkipProof.make_thm thy (Var (("A", 0), propT))) i) st |
|
118 |
||
31550
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31549
diff
changeset
|
119 |
fun remove_last_goal thy st = |
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31549
diff
changeset
|
120 |
(Tactic.rtac (SkipProof.make_thm thy (Var (("A", 0), propT))) (nprems_of st)) st |
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31549
diff
changeset
|
121 |
|
31514 | 122 |
(* reference to preprocessing of InductiveSet package *) |
123 |
||
31723
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
31580
diff
changeset
|
124 |
val ind_set_codegen_preproc = Inductive_Set.codegen_preproc; |
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
125 |
|
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
126 |
(** fundamentals **) |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
127 |
|
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
128 |
(* syntactic operations *) |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
129 |
|
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
130 |
fun mk_eq (x, xs) = |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
131 |
let fun mk_eqs _ [] = [] |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
132 |
| mk_eqs a (b::cs) = |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
133 |
HOLogic.mk_eq (Free (a, fastype_of b), b) :: mk_eqs a cs |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
134 |
in mk_eqs x xs end; |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
135 |
|
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
136 |
fun mk_tupleT [] = HOLogic.unitT |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
137 |
| mk_tupleT Ts = foldr1 HOLogic.mk_prodT Ts; |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
138 |
|
31514 | 139 |
fun dest_tupleT (Type (@{type_name Product_Type.unit}, [])) = [] |
140 |
| dest_tupleT (Type (@{type_name "*"}, [T1, T2])) = T1 :: (dest_tupleT T2) |
|
141 |
| dest_tupleT t = [t] |
|
142 |
||
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
143 |
fun mk_tuple [] = HOLogic.unit |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
144 |
| mk_tuple ts = foldr1 HOLogic.mk_prod ts; |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
145 |
|
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
146 |
fun dest_tuple (Const (@{const_name Product_Type.Unity}, _)) = [] |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
147 |
| dest_tuple (Const (@{const_name Pair}, _) $ t1 $ t2) = t1 :: (dest_tuple t2) |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
148 |
| dest_tuple t = [t] |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
149 |
|
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
150 |
fun mk_scomp (t, u) = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
151 |
let |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
152 |
val T = fastype_of t |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
153 |
val U = fastype_of u |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
154 |
val [A] = binder_types T |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
155 |
val D = body_type U |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
156 |
in |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
157 |
Const (@{const_name "scomp"}, T --> U --> A --> D) $ t $ u |
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
158 |
end; |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
159 |
|
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
160 |
fun dest_funT (Type ("fun",[S, T])) = (S, T) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
161 |
| dest_funT T = raise TYPE ("dest_funT", [T], []) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
162 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
163 |
fun mk_fun_comp (t, u) = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
164 |
let |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
165 |
val (_, B) = dest_funT (fastype_of t) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
166 |
val (C, A) = dest_funT (fastype_of u) |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
167 |
in |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
168 |
Const(@{const_name "Fun.comp"}, (A --> B) --> (C --> A) --> C --> B) $ t $ u |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
169 |
end; |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
170 |
|
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
171 |
fun dest_randomT (Type ("fun", [@{typ Random.seed}, |
32657 | 172 |
Type ("*", [Type ("*", [T, @{typ "unit => Code_Evaluation.term"}]) ,@{typ Random.seed}])])) = T |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
173 |
| dest_randomT T = raise TYPE ("dest_randomT", [T], []) |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
174 |
|
31514 | 175 |
(* destruction of intro rules *) |
176 |
||
177 |
(* FIXME: look for other place where this functionality was used before *) |
|
178 |
fun strip_intro_concl nparams intro = let |
|
179 |
val _ $ u = Logic.strip_imp_concl intro |
|
180 |
val (pred, all_args) = strip_comb u |
|
181 |
val (params, args) = chop nparams all_args |
|
182 |
in (pred, (params, args)) end |
|
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
183 |
|
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
184 |
(** data structures **) |
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
185 |
|
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
186 |
type smode = int list; |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
187 |
type mode = smode option list * smode; |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
188 |
datatype tmode = Mode of mode * int list * tmode option list; |
31514 | 189 |
|
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
190 |
fun split_smode is ts = |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
191 |
let |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
192 |
fun split_smode' _ _ [] = ([], []) |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
193 |
| split_smode' is i (t::ts) = (if i mem is then apfst else apsnd) (cons t) |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
194 |
(split_smode' is (i+1) ts) |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
195 |
in split_smode' is 1 ts end |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
196 |
|
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
197 |
fun split_mode (iss, is) ts = |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
198 |
let |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
199 |
val (t1, t2) = chop (length iss) ts |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
200 |
in (t1, split_smode is t2) end |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
201 |
|
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
202 |
fun string_of_mode (iss, is) = space_implode " -> " (map |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
203 |
(fn NONE => "X" |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
204 |
| SOME js => enclose "[" "]" (commas (map string_of_int js))) |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
205 |
(iss @ [SOME is])); |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
206 |
|
32316
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
207 |
fun string_of_tmode (Mode (predmode, termmode, param_modes)) = |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
208 |
"predmode: " ^ (string_of_mode predmode) ^ |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
209 |
(if null param_modes then "" else |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
210 |
"; " ^ "params: " ^ commas (map (the_default "NONE" o Option.map string_of_tmode) param_modes)) |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
211 |
|
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
212 |
datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term | |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
213 |
GeneratorPrem of term list * term | Generator of (string * typ); |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
214 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
215 |
type moded_clause = term list * (indprem * tmode) list |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
216 |
type 'a pred_mode_table = (string * (mode * 'a) list) list |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
217 |
|
31514 | 218 |
datatype predfun_data = PredfunData of { |
219 |
name : string, |
|
220 |
definition : thm, |
|
221 |
intro : thm, |
|
222 |
elim : thm |
|
223 |
}; |
|
224 |
||
225 |
fun rep_predfun_data (PredfunData data) = data; |
|
226 |
fun mk_predfun_data (name, definition, intro, elim) = |
|
227 |
PredfunData {name = name, definition = definition, intro = intro, elim = elim} |
|
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
228 |
|
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
229 |
datatype function_data = FunctionData of { |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
230 |
name : string, |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
231 |
equation : thm option (* is not used at all? *) |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
232 |
}; |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
233 |
|
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
234 |
fun rep_function_data (FunctionData data) = data; |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
235 |
fun mk_function_data (name, equation) = |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
236 |
FunctionData {name = name, equation = equation} |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
237 |
|
31514 | 238 |
datatype pred_data = PredData of { |
239 |
intros : thm list, |
|
240 |
elim : thm option, |
|
241 |
nparams : int, |
|
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
242 |
functions : (mode * predfun_data) list, |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
243 |
generators : (mode * function_data) list, |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
244 |
sizelim_functions : (mode * function_data) list |
31514 | 245 |
}; |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
246 |
|
31514 | 247 |
fun rep_pred_data (PredData data) = data; |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
248 |
fun mk_pred_data ((intros, elim, nparams), (functions, generators, sizelim_functions)) = |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
249 |
PredData {intros = intros, elim = elim, nparams = nparams, |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
250 |
functions = functions, generators = generators, sizelim_functions = sizelim_functions} |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
251 |
fun map_pred_data f (PredData {intros, elim, nparams, functions, generators, sizelim_functions}) = |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
252 |
mk_pred_data (f ((intros, elim, nparams), (functions, generators, sizelim_functions))) |
31514 | 253 |
|
254 |
fun eq_option eq (NONE, NONE) = true |
|
255 |
| eq_option eq (SOME x, SOME y) = eq (x, y) |
|
256 |
| eq_option eq _ = false |
|
257 |
||
258 |
fun eq_pred_data (PredData d1, PredData d2) = |
|
259 |
eq_list (Thm.eq_thm) (#intros d1, #intros d2) andalso |
|
260 |
eq_option (Thm.eq_thm) (#elim d1, #elim d2) andalso |
|
261 |
#nparams d1 = #nparams d2 |
|
262 |
||
263 |
structure PredData = TheoryDataFun |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
264 |
( |
31514 | 265 |
type T = pred_data Graph.T; |
266 |
val empty = Graph.empty; |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
267 |
val copy = I; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
268 |
val extend = I; |
31514 | 269 |
fun merge _ = Graph.merge eq_pred_data; |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
270 |
); |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
271 |
|
31514 | 272 |
(* queries *) |
273 |
||
31549 | 274 |
fun lookup_pred_data thy name = |
31573
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
275 |
Option.map rep_pred_data (try (Graph.get_node (PredData.get thy)) name) |
31514 | 276 |
|
277 |
fun the_pred_data thy name = case lookup_pred_data thy name |
|
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
278 |
of NONE => error ("No such predicate " ^ quote name) |
31514 | 279 |
| SOME data => data; |
280 |
||
31877
e3de75d3b898
exported is_registered; added debug messages; removed extended parameter compilation in predicate compiler
bulwahn
parents:
31876
diff
changeset
|
281 |
val is_registered = is_some oo lookup_pred_data |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
282 |
|
31514 | 283 |
val all_preds_of = Graph.keys o PredData.get |
284 |
||
285 |
val intros_of = #intros oo the_pred_data |
|
286 |
||
287 |
fun the_elim_of thy name = case #elim (the_pred_data thy name) |
|
288 |
of NONE => error ("No elimination rule for predicate " ^ quote name) |
|
289 |
| SOME thm => thm |
|
290 |
||
291 |
val has_elim = is_some o #elim oo the_pred_data; |
|
292 |
||
293 |
val nparams_of = #nparams oo the_pred_data |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
294 |
|
31514 | 295 |
val modes_of = (map fst) o #functions oo the_pred_data |
296 |
||
297 |
fun all_modes_of thy = map (fn name => (name, modes_of thy name)) (all_preds_of thy) |
|
298 |
||
299 |
val is_compiled = not o null o #functions oo the_pred_data |
|
300 |
||
301 |
fun lookup_predfun_data thy name mode = |
|
302 |
Option.map rep_predfun_data (AList.lookup (op =) |
|
303 |
(#functions (the_pred_data thy name)) mode) |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
304 |
|
31514 | 305 |
fun the_predfun_data thy name mode = case lookup_predfun_data thy name mode |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
306 |
of NONE => error ("No function defined for mode " ^ string_of_mode mode ^ " of predicate " ^ name) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
307 |
| SOME data => data; |
31514 | 308 |
|
309 |
val predfun_name_of = #name ooo the_predfun_data |
|
310 |
||
311 |
val predfun_definition_of = #definition ooo the_predfun_data |
|
312 |
||
313 |
val predfun_intro_of = #intro ooo the_predfun_data |
|
314 |
||
315 |
val predfun_elim_of = #elim ooo the_predfun_data |
|
316 |
||
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
317 |
fun lookup_generator_data thy name mode = |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
318 |
Option.map rep_function_data (AList.lookup (op =) |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
319 |
(#generators (the_pred_data thy name)) mode) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
320 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
321 |
fun the_generator_data thy name mode = case lookup_generator_data thy name mode |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
322 |
of NONE => error ("No generator defined for mode " ^ string_of_mode mode ^ " of predicate " ^ name) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
323 |
| SOME data => data |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
324 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
325 |
val generator_name_of = #name ooo the_generator_data |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
326 |
|
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
327 |
val generator_modes_of = (map fst) o #generators oo the_pred_data |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
328 |
|
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
329 |
fun all_generator_modes_of thy = |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
330 |
map (fn name => (name, generator_modes_of thy name)) (all_preds_of thy) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
331 |
|
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
332 |
fun lookup_sizelim_function_data thy name mode = |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
333 |
Option.map rep_function_data (AList.lookup (op =) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
334 |
(#sizelim_functions (the_pred_data thy name)) mode) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
335 |
|
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
336 |
fun the_sizelim_function_data thy name mode = case lookup_sizelim_function_data thy name mode |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
337 |
of NONE => error ("No size-limited function defined for mode " ^ string_of_mode mode |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
338 |
^ " of predicate " ^ name) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
339 |
| SOME data => data |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
340 |
|
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
341 |
val sizelim_function_name_of = #name ooo the_sizelim_function_data |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
342 |
|
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
343 |
(*val generator_modes_of = (map fst) o #generators oo the_pred_data*) |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
344 |
|
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
345 |
(* diagnostic display functions *) |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
346 |
|
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
347 |
fun print_modes modes = Output.tracing ("Inferred modes:\n" ^ |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
348 |
cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
349 |
string_of_mode ms)) modes)); |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
350 |
|
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
351 |
fun print_pred_mode_table string_of_entry thy pred_mode_table = |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
352 |
let |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
353 |
fun print_mode pred (mode, entry) = "mode : " ^ (string_of_mode mode) |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
354 |
^ (string_of_entry pred mode entry) |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
355 |
fun print_pred (pred, modes) = |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
356 |
"predicate " ^ pred ^ ": " ^ cat_lines (map (print_mode pred) modes) |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
357 |
val _ = Output.tracing (cat_lines (map print_pred pred_mode_table)) |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
358 |
in () end; |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
359 |
|
32316
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
360 |
fun string_of_moded_prem thy (Prem (ts, p), tmode) = |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
361 |
(Syntax.string_of_term_global thy (list_comb (p, ts))) ^ |
32316
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
362 |
"(" ^ (string_of_tmode tmode) ^ ")" |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
363 |
| string_of_moded_prem thy (GeneratorPrem (ts, p), Mode (predmode, is, _)) = |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
364 |
(Syntax.string_of_term_global thy (list_comb (p, ts))) ^ |
32316
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
365 |
"(generator_mode: " ^ (string_of_mode predmode) ^ ")" |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
366 |
| string_of_moded_prem thy (Generator (v, T), _) = |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
367 |
"Generator for " ^ v ^ " of Type " ^ (Syntax.string_of_typ_global thy T) |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
368 |
| string_of_moded_prem thy (Negprem (ts, p), Mode (_, is, _)) = |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
369 |
(Syntax.string_of_term_global thy (list_comb (p, ts))) ^ |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
370 |
"(negative mode: " ^ (space_implode ", " (map string_of_int is)) ^ ")" |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
371 |
| string_of_moded_prem thy (Sidecond t, Mode (_, is, _)) = |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
372 |
(Syntax.string_of_term_global thy t) ^ |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
373 |
"(sidecond mode: " ^ (space_implode ", " (map string_of_int is)) ^ ")" |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
374 |
| string_of_moded_prem _ _ = error "string_of_moded_prem: unimplemented" |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
375 |
|
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
376 |
fun print_moded_clauses thy = |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
377 |
let |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
378 |
fun string_of_clause pred mode clauses = |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
379 |
cat_lines (map (fn (ts, prems) => (space_implode " --> " |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
380 |
(map (string_of_moded_prem thy) prems)) ^ " --> " ^ pred ^ " " |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
381 |
^ (space_implode " " (map (Syntax.string_of_term_global thy) ts))) clauses) |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
382 |
in print_pred_mode_table string_of_clause thy end; |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
383 |
|
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
384 |
fun print_compiled_terms thy = |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
385 |
print_pred_mode_table (fn _ => fn _ => Syntax.string_of_term_global thy) thy |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
386 |
|
31876
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
387 |
fun print_stored_rules thy = |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
388 |
let |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
389 |
val preds = (Graph.keys o PredData.get) thy |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
390 |
fun print pred () = let |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
391 |
val _ = writeln ("predicate: " ^ pred) |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
392 |
val _ = writeln ("number of parameters: " ^ string_of_int (nparams_of thy pred)) |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
393 |
val _ = writeln ("introrules: ") |
32091
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents:
31986
diff
changeset
|
394 |
val _ = fold (fn thm => fn u => writeln (Display.string_of_thm_global thy thm)) |
31876
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
395 |
(rev (intros_of thy pred)) () |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
396 |
in |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
397 |
if (has_elim thy pred) then |
32091
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents:
31986
diff
changeset
|
398 |
writeln ("elimrule: " ^ Display.string_of_thm_global thy (the_elim_of thy pred)) |
31876
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
399 |
else |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
400 |
writeln ("no elimrule defined") |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
401 |
end |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
402 |
in |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
403 |
fold print preds () |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
404 |
end; |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
405 |
|
32306
19f55947d4d5
removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents:
32287
diff
changeset
|
406 |
fun print_all_modes thy = |
19f55947d4d5
removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents:
32287
diff
changeset
|
407 |
let |
19f55947d4d5
removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents:
32287
diff
changeset
|
408 |
val _ = writeln ("Inferred modes:") |
19f55947d4d5
removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents:
32287
diff
changeset
|
409 |
fun print (pred, modes) u = |
19f55947d4d5
removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents:
32287
diff
changeset
|
410 |
let |
19f55947d4d5
removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents:
32287
diff
changeset
|
411 |
val _ = writeln ("predicate: " ^ pred) |
19f55947d4d5
removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents:
32287
diff
changeset
|
412 |
val _ = writeln ("modes: " ^ (commas (map string_of_mode modes))) |
19f55947d4d5
removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents:
32287
diff
changeset
|
413 |
in u end |
19f55947d4d5
removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents:
32287
diff
changeset
|
414 |
in |
19f55947d4d5
removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents:
32287
diff
changeset
|
415 |
fold print (all_modes_of thy) () |
19f55947d4d5
removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents:
32287
diff
changeset
|
416 |
end |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
417 |
|
31876
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
418 |
(** preprocessing rules **) |
31573
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
419 |
|
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
420 |
fun imp_prems_conv cv ct = |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
421 |
case Thm.term_of ct of |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
422 |
Const ("==>", _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv) (imp_prems_conv cv) ct |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
423 |
| _ => Conv.all_conv ct |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
424 |
|
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
425 |
fun Trueprop_conv cv ct = |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
426 |
case Thm.term_of ct of |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
427 |
Const ("Trueprop", _) $ _ => Conv.arg_conv cv ct |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
428 |
| _ => error "Trueprop_conv" |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
429 |
|
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
430 |
fun preprocess_intro thy rule = |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
431 |
Conv.fconv_rule |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
432 |
(imp_prems_conv |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
433 |
(Trueprop_conv (Conv.try_conv (Conv.rewr_conv (Thm.symmetric @{thm Predicate.eq_is_eq}))))) |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
434 |
(Thm.transfer thy rule) |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
435 |
|
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
436 |
fun preprocess_elim thy nparams elimrule = |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
437 |
let |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
438 |
fun replace_eqs (Const ("Trueprop", _) $ (Const ("op =", T) $ lhs $ rhs)) = |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
439 |
HOLogic.mk_Trueprop (Const (@{const_name Predicate.eq}, T) $ lhs $ rhs) |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
440 |
| replace_eqs t = t |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
441 |
val prems = Thm.prems_of elimrule |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
442 |
val nargs = length (snd (strip_comb (HOLogic.dest_Trueprop (hd prems)))) - nparams |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
443 |
fun preprocess_case t = |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
444 |
let |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
445 |
val params = Logic.strip_params t |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
446 |
val (assums1, assums2) = chop nargs (Logic.strip_assums_hyp t) |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
447 |
val assums_hyp' = assums1 @ (map replace_eqs assums2) |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
448 |
in |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
449 |
list_all (params, Logic.list_implies (assums_hyp', Logic.strip_assums_concl t)) |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
450 |
end |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
451 |
val cases' = map preprocess_case (tl prems) |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
452 |
val elimrule' = Logic.list_implies ((hd prems) :: cases', Thm.concl_of elimrule) |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
453 |
in |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
454 |
Thm.equal_elim |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
455 |
(Thm.symmetric (Conv.implies_concl_conv (MetaSimplifier.rewrite true [@{thm eq_is_eq}]) |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
456 |
(cterm_of thy elimrule'))) |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
457 |
elimrule |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
458 |
end; |
31573
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
459 |
|
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
460 |
(* special case: predicate with no introduction rule *) |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
461 |
fun noclause thy predname elim = let |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
462 |
val T = (Logic.unvarifyT o Sign.the_const_type thy) predname |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
463 |
val Ts = binder_types T |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
464 |
val names = Name.variant_list [] |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
465 |
(map (fn i => "x" ^ (string_of_int i)) (1 upto (length Ts))) |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
466 |
val vs = map2 (curry Free) names Ts |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
467 |
val clausehd = HOLogic.mk_Trueprop (list_comb (Const (predname, T), vs)) |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
468 |
val intro_t = Logic.mk_implies (@{prop False}, clausehd) |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
469 |
val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT)) |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
470 |
val elim_t = Logic.list_implies ([clausehd, Logic.mk_implies (@{prop False}, P)], P) |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
471 |
val intro = Goal.prove (ProofContext.init thy) names [] intro_t |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
472 |
(fn {...} => etac @{thm FalseE} 1) |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
473 |
val elim = Goal.prove (ProofContext.init thy) ("P" :: names) [] elim_t |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
474 |
(fn {...} => etac elim 1) |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
475 |
in |
32314
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
476 |
([intro], elim) |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
477 |
end |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
478 |
|
31573
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
479 |
fun fetch_pred_data thy name = |
31723
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
31580
diff
changeset
|
480 |
case try (Inductive.the_inductive (ProofContext.init thy)) name of |
31573
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
481 |
SOME (info as (_, result)) => |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
482 |
let |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
483 |
fun is_intro_of intro = |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
484 |
let |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
485 |
val (const, _) = strip_comb (HOLogic.dest_Trueprop (concl_of intro)) |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
486 |
in (fst (dest_Const const) = name) end; |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
487 |
val intros = ind_set_codegen_preproc thy ((map (preprocess_intro thy)) |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
488 |
(filter is_intro_of (#intrs result))) |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
489 |
val pre_elim = nth (#elims result) (find_index (fn s => s = name) (#names (fst info))) |
31723
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
31580
diff
changeset
|
490 |
val nparams = length (Inductive.params_of (#raw_induct result)) |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
491 |
val elim = singleton (ind_set_codegen_preproc thy) (preprocess_elim thy nparams pre_elim) |
32314
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
492 |
val (intros, elim) = if null intros then noclause thy name elim else (intros, elim) |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
493 |
in |
32314
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
494 |
mk_pred_data ((intros, SOME elim, nparams), ([], [], [])) |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
495 |
end |
31573
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
496 |
| NONE => error ("No such predicate: " ^ quote name) |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
497 |
|
31514 | 498 |
(* updaters *) |
499 |
||
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
500 |
fun apfst3 f (x, y, z) = (f x, y, z) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
501 |
fun apsnd3 f (x, y, z) = (x, f y, z) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
502 |
fun aptrd3 f (x, y, z) = (x, y, f z) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
503 |
|
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
504 |
fun add_predfun name mode data = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
505 |
let |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
506 |
val add = (apsnd o apfst3 o cons) (mode, mk_predfun_data data) |
31514 | 507 |
in PredData.map (Graph.map_node name (map_pred_data add)) end |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
508 |
|
31876
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
509 |
fun is_inductive_predicate thy name = |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
510 |
is_some (try (Inductive.the_inductive (ProofContext.init thy)) name) |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
511 |
|
32314
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
512 |
fun depending_preds_of thy (key, value) = |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
513 |
let |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
514 |
val intros = (#intros o rep_pred_data) value |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
515 |
in |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
516 |
fold Term.add_const_names (map Thm.prop_of intros) [] |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
517 |
|> filter (fn c => (not (c = key)) andalso (is_inductive_predicate thy c orelse is_registered thy c)) |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
518 |
end; |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
519 |
|
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
520 |
|
31876
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
521 |
(* code dependency graph *) |
32314
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
522 |
(* |
31876
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
523 |
fun dependencies_of thy name = |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
524 |
let |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
525 |
val (intros, elim, nparams) = fetch_pred_data thy name |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
526 |
val data = mk_pred_data ((intros, SOME elim, nparams), ([], [], [])) |
31876
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
527 |
val keys = depending_preds_of thy intros |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
528 |
in |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
529 |
(data, keys) |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
530 |
end; |
32314
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
531 |
*) |
32315
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
532 |
(* guessing number of parameters *) |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
533 |
fun find_indexes pred xs = |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
534 |
let |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
535 |
fun find is n [] = is |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
536 |
| find is n (x :: xs) = find (if pred x then (n :: is) else is) (n + 1) xs; |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
537 |
in rev (find [] 0 xs) end; |
31876
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
538 |
|
32315
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
539 |
fun is_predT (T as Type("fun", [_, _])) = (snd (strip_type T) = HOLogic.boolT) |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
540 |
| is_predT _ = false |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
541 |
|
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
542 |
fun guess_nparams T = |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
543 |
let |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
544 |
val argTs = binder_types T |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
545 |
val nparams = fold (curry Int.max) |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
546 |
(map (fn x => x + 1) (find_indexes is_predT argTs)) 0 |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
547 |
in nparams end; |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
548 |
|
31573
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
549 |
fun add_intro thm thy = let |
32315
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
550 |
val (name, T) = dest_Const (fst (strip_intro_concl 0 (prop_of thm))) |
31573
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
551 |
fun cons_intro gr = |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
552 |
case try (Graph.get_node gr) name of |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
553 |
SOME pred_data => Graph.map_node name (map_pred_data |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
554 |
(apfst (fn (intro, elim, nparams) => (thm::intro, elim, nparams)))) gr |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
555 |
| NONE => |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
556 |
let |
32315
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
557 |
val nparams = the_default (guess_nparams T) (try (#nparams o rep_pred_data o (fetch_pred_data thy)) name) |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
558 |
in Graph.new_node (name, mk_pred_data (([thm], NONE, nparams), ([], [], []))) gr end; |
31573
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
559 |
in PredData.map cons_intro thy end |
31514 | 560 |
|
561 |
fun set_elim thm = let |
|
562 |
val (name, _) = dest_Const (fst |
|
563 |
(strip_comb (HOLogic.dest_Trueprop (hd (prems_of thm))))) |
|
564 |
fun set (intros, _, nparams) = (intros, SOME thm, nparams) |
|
565 |
in PredData.map (Graph.map_node name (map_pred_data (apfst set))) end |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
566 |
|
31514 | 567 |
fun set_nparams name nparams = let |
568 |
fun set (intros, elim, _ ) = (intros, elim, nparams) |
|
569 |
in PredData.map (Graph.map_node name (map_pred_data (apfst set))) end |
|
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
570 |
|
32314
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
571 |
fun register_predicate (pre_intros, pre_elim, nparams) thy = let |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
572 |
val (name, _) = dest_Const (fst (strip_intro_concl nparams (prop_of (hd pre_intros)))) |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
573 |
(* preprocessing *) |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
574 |
val intros = ind_set_codegen_preproc thy (map (preprocess_intro thy) pre_intros) |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
575 |
val elim = singleton (ind_set_codegen_preproc thy) (preprocess_elim thy nparams pre_elim) |
31876
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
576 |
in |
32314
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
577 |
PredData.map |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
578 |
(Graph.new_node (name, mk_pred_data ((intros, SOME elim, nparams), ([], [], [])))) thy |
31876
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
579 |
end |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
580 |
|
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
581 |
fun set_generator_name pred mode name = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
582 |
let |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
583 |
val set = (apsnd o apsnd3 o cons) (mode, mk_function_data (name, NONE)) |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
584 |
in |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
585 |
PredData.map (Graph.map_node pred (map_pred_data set)) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
586 |
end |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
587 |
|
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
588 |
fun set_sizelim_function_name pred mode name = |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
589 |
let |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
590 |
val set = (apsnd o aptrd3 o cons) (mode, mk_function_data (name, NONE)) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
591 |
in |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
592 |
PredData.map (Graph.map_node pred (map_pred_data set)) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
593 |
end |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
594 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
595 |
(** data structures for generic compilation for different monads **) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
596 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
597 |
(* maybe rename functions more generic: |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
598 |
mk_predT -> mk_monadT; dest_predT -> dest_monadT |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
599 |
mk_single -> mk_return (?) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
600 |
*) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
601 |
datatype compilation_funs = CompilationFuns of { |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
602 |
mk_predT : typ -> typ, |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
603 |
dest_predT : typ -> typ, |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
604 |
mk_bot : typ -> term, |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
605 |
mk_single : term -> term, |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
606 |
mk_bind : term * term -> term, |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
607 |
mk_sup : term * term -> term, |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
608 |
mk_if : term -> term, |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
609 |
mk_not : term -> term, |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
610 |
(* funT_of : mode -> typ -> typ, *) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
611 |
(* mk_fun_of : theory -> (string * typ) -> mode -> term, *) |
32351 | 612 |
mk_map : typ -> typ -> term -> term -> term, |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
613 |
lift_pred : term -> term |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
614 |
}; |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
615 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
616 |
fun mk_predT (CompilationFuns funs) = #mk_predT funs |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
617 |
fun dest_predT (CompilationFuns funs) = #dest_predT funs |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
618 |
fun mk_bot (CompilationFuns funs) = #mk_bot funs |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
619 |
fun mk_single (CompilationFuns funs) = #mk_single funs |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
620 |
fun mk_bind (CompilationFuns funs) = #mk_bind funs |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
621 |
fun mk_sup (CompilationFuns funs) = #mk_sup funs |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
622 |
fun mk_if (CompilationFuns funs) = #mk_if funs |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
623 |
fun mk_not (CompilationFuns funs) = #mk_not funs |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
624 |
(*fun funT_of (CompilationFuns funs) = #funT_of funs*) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
625 |
(*fun mk_fun_of (CompilationFuns funs) = #mk_fun_of funs*) |
32351 | 626 |
fun mk_map (CompilationFuns funs) = #mk_map funs |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
627 |
fun lift_pred (CompilationFuns funs) = #lift_pred funs |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
628 |
|
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
629 |
fun funT_of compfuns (iss, is) T = |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
630 |
let |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
631 |
val Ts = binder_types T |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
632 |
val (paramTs, (inargTs, outargTs)) = split_mode (iss, is) Ts |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
633 |
val paramTs' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) iss paramTs |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
634 |
in |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
635 |
(paramTs' @ inargTs) ---> (mk_predT compfuns (mk_tupleT outargTs)) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
636 |
end; |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
637 |
|
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
638 |
fun sizelim_funT_of compfuns (iss, is) T = |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
639 |
let |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
640 |
val Ts = binder_types T |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
641 |
val (paramTs, (inargTs, outargTs)) = split_mode (iss, is) Ts |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
642 |
val paramTs' = map2 (fn SOME is => sizelim_funT_of compfuns ([], is) | NONE => I) iss paramTs |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
643 |
in |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
644 |
(paramTs' @ inargTs @ [@{typ "code_numeral"}]) ---> (mk_predT compfuns (mk_tupleT outargTs)) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
645 |
end; |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
646 |
|
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
647 |
fun mk_fun_of compfuns thy (name, T) mode = |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
648 |
Const (predfun_name_of thy name mode, funT_of compfuns mode T) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
649 |
|
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
650 |
fun mk_sizelim_fun_of compfuns thy (name, T) mode = |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
651 |
Const (sizelim_function_name_of thy name mode, sizelim_funT_of compfuns mode T) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
652 |
|
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
653 |
fun mk_generator_of compfuns thy (name, T) mode = |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
654 |
Const (generator_name_of thy name mode, sizelim_funT_of compfuns mode T) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
655 |
|
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
656 |
|
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
657 |
structure PredicateCompFuns = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
658 |
struct |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
659 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
660 |
fun mk_predT T = Type (@{type_name "Predicate.pred"}, [T]) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
661 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
662 |
fun dest_predT (Type (@{type_name "Predicate.pred"}, [T])) = T |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
663 |
| dest_predT T = raise TYPE ("dest_predT", [T], []); |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
664 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
665 |
fun mk_bot T = Const (@{const_name Orderings.bot}, mk_predT T); |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
666 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
667 |
fun mk_single t = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
668 |
let val T = fastype_of t |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
669 |
in Const(@{const_name Predicate.single}, T --> mk_predT T) $ t end; |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
670 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
671 |
fun mk_bind (x, f) = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
672 |
let val T as Type ("fun", [_, U]) = fastype_of f |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
673 |
in |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
674 |
Const (@{const_name Predicate.bind}, fastype_of x --> T --> U) $ x $ f |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
675 |
end; |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
676 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
677 |
val mk_sup = HOLogic.mk_binop @{const_name sup}; |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
678 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
679 |
fun mk_if cond = Const (@{const_name Predicate.if_pred}, |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
680 |
HOLogic.boolT --> mk_predT HOLogic.unitT) $ cond; |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
681 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
682 |
fun mk_not t = let val T = mk_predT HOLogic.unitT |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
683 |
in Const (@{const_name Predicate.not_pred}, T --> T) $ t end |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
684 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
685 |
fun mk_Enum f = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
686 |
let val T as Type ("fun", [T', _]) = fastype_of f |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
687 |
in |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
688 |
Const (@{const_name Predicate.Pred}, T --> mk_predT T') $ f |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
689 |
end; |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
690 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
691 |
fun mk_Eval (f, x) = |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
692 |
let |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
693 |
val T = fastype_of x |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
694 |
in |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
695 |
Const (@{const_name Predicate.eval}, mk_predT T --> T --> HOLogic.boolT) $ f $ x |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
696 |
end; |
32351 | 697 |
|
698 |
fun mk_map T1 T2 tf tp = Const (@{const_name Predicate.map}, |
|
699 |
(T1 --> T2) --> mk_predT T1 --> mk_predT T2) $ tf $ tp; |
|
700 |
||
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
701 |
val lift_pred = I |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
702 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
703 |
val compfuns = CompilationFuns {mk_predT = mk_predT, dest_predT = dest_predT, mk_bot = mk_bot, |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
704 |
mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if, mk_not = mk_not, |
32351 | 705 |
mk_map = mk_map, lift_pred = lift_pred}; |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
706 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
707 |
end; |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
708 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
709 |
(* termify_code: |
32657 | 710 |
val termT = Type ("Code_Evaluation.term", []); |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
711 |
fun termifyT T = HOLogic.mk_prodT (T, HOLogic.unitT --> termT) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
712 |
*) |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
713 |
(* |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
714 |
fun lift_random random = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
715 |
let |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
716 |
val T = dest_randomT (fastype_of random) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
717 |
in |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
718 |
mk_scomp (random, |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
719 |
mk_fun_comp (HOLogic.pair_const (PredicateCompFuns.mk_predT T) @{typ Random.seed}, |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
720 |
mk_fun_comp (Const (@{const_name Predicate.single}, T --> (PredicateCompFuns.mk_predT T)), |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
721 |
Const (@{const_name "fst"}, HOLogic.mk_prodT (T, @{typ "unit => term"}) --> T)))) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
722 |
end; |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
723 |
*) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
724 |
|
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
725 |
structure RPredCompFuns = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
726 |
struct |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
727 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
728 |
fun mk_rpredT T = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
729 |
@{typ "Random.seed"} --> HOLogic.mk_prodT (PredicateCompFuns.mk_predT T, @{typ "Random.seed"}) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
730 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
731 |
fun dest_rpredT (Type ("fun", [_, |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
732 |
Type (@{type_name "*"}, [Type (@{type_name "Predicate.pred"}, [T]), _])])) = T |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
733 |
| dest_rpredT T = raise TYPE ("dest_rpredT", [T], []); |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
734 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
735 |
fun mk_bot T = Const(@{const_name RPred.bot}, mk_rpredT T) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
736 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
737 |
fun mk_single t = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
738 |
let |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
739 |
val T = fastype_of t |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
740 |
in |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
741 |
Const (@{const_name RPred.return}, T --> mk_rpredT T) $ t |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
742 |
end; |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
743 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
744 |
fun mk_bind (x, f) = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
745 |
let |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
746 |
val T as (Type ("fun", [_, U])) = fastype_of f |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
747 |
in |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
748 |
Const (@{const_name RPred.bind}, fastype_of x --> T --> U) $ x $ f |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
749 |
end |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
750 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
751 |
val mk_sup = HOLogic.mk_binop @{const_name RPred.supp} |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
752 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
753 |
fun mk_if cond = Const (@{const_name RPred.if_rpred}, |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
754 |
HOLogic.boolT --> mk_rpredT HOLogic.unitT) $ cond; |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
755 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
756 |
fun mk_not t = error "Negation is not defined for RPred" |
32351 | 757 |
|
758 |
fun mk_map t = error "FIXME" (*FIXME*) |
|
759 |
||
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
760 |
fun lift_pred t = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
761 |
let |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
762 |
val T = PredicateCompFuns.dest_predT (fastype_of t) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
763 |
val lift_predT = PredicateCompFuns.mk_predT T --> mk_rpredT T |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
764 |
in |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
765 |
Const (@{const_name "RPred.lift_pred"}, lift_predT) $ t |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
766 |
end; |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
767 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
768 |
val compfuns = CompilationFuns {mk_predT = mk_rpredT, dest_predT = dest_rpredT, mk_bot = mk_bot, |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
769 |
mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if, mk_not = mk_not, |
32351 | 770 |
mk_map = mk_map, lift_pred = lift_pred}; |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
771 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
772 |
end; |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
773 |
(* for external use with interactive mode *) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
774 |
val rpred_compfuns = RPredCompFuns.compfuns; |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
775 |
|
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
776 |
fun lift_random random = |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
777 |
let |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
778 |
val T = dest_randomT (fastype_of random) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
779 |
in |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
780 |
Const (@{const_name lift_random}, (@{typ Random.seed} --> |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
781 |
HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed})) --> |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
782 |
RPredCompFuns.mk_rpredT T) $ random |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
783 |
end; |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
784 |
|
31514 | 785 |
(* Mode analysis *) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
786 |
|
31514 | 787 |
(*** check if a term contains only constructor functions ***) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
788 |
fun is_constrt thy = |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
789 |
let |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
790 |
val cnstrs = flat (maps |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
791 |
(map (fn (_, (Tname, _, cs)) => map (apsnd (rpair Tname o length)) cs) o #descr o snd) |
31784 | 792 |
(Symtab.dest (Datatype.get_all thy))); |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
793 |
fun check t = (case strip_comb t of |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
794 |
(Free _, []) => true |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
795 |
| (Const (s, T), ts) => (case (AList.lookup (op =) cnstrs s, body_type T) of |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
796 |
(SOME (i, Tname), Type (Tname', _)) => length ts = i andalso Tname = Tname' andalso forall check ts |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
797 |
| _ => false) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
798 |
| _ => false) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
799 |
in check end; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
800 |
|
31514 | 801 |
(*** check if a type is an equality type (i.e. doesn't contain fun) |
802 |
FIXME this is only an approximation ***) |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
803 |
fun is_eqT (Type (s, Ts)) = s <> "fun" andalso forall is_eqT Ts |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
804 |
| is_eqT _ = true; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
805 |
|
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
806 |
fun term_vs tm = fold_aterms (fn Free (x, T) => cons x | _ => I) tm []; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
807 |
val terms_vs = distinct (op =) o maps term_vs; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
808 |
|
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
809 |
(** collect all Frees in a term (with duplicates!) **) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
810 |
fun term_vTs tm = |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
811 |
fold_aterms (fn Free xT => cons xT | _ => I) tm []; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
812 |
|
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
813 |
(*FIXME this function should not be named merge... make it local instead*) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
814 |
fun merge xs [] = xs |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
815 |
| merge [] ys = ys |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
816 |
| merge (x::xs) (y::ys) = if length x >= length y then x::merge xs (y::ys) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
817 |
else y::merge (x::xs) ys; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
818 |
|
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
819 |
fun subsets i j = if i <= j then |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
820 |
let val is = subsets (i+1) j |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
821 |
in merge (map (fn ks => i::ks) is) is end |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
822 |
else [[]]; |
31514 | 823 |
|
824 |
(* FIXME: should be in library - map_prod *) |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
825 |
fun cprod ([], ys) = [] |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
826 |
| cprod (x :: xs, ys) = map (pair x) ys @ cprod (xs, ys); |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
827 |
|
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
828 |
fun cprods xss = foldr (map op :: o cprod) [[]] xss; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
829 |
|
31514 | 830 |
|
31515 | 831 |
|
31514 | 832 |
(*TODO: cleanup function and put together with modes_of_term *) |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
833 |
(* |
31514 | 834 |
fun modes_of_param default modes t = let |
835 |
val (vs, t') = strip_abs t |
|
836 |
val b = length vs |
|
837 |
fun mk_modes name args = Option.map (maps (fn (m as (iss, is)) => |
|
838 |
let |
|
839 |
val (args1, args2) = |
|
840 |
if length args < length iss then |
|
841 |
error ("Too few arguments for inductive predicate " ^ name) |
|
842 |
else chop (length iss) args; |
|
843 |
val k = length args2; |
|
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
844 |
val perm = map (fn i => (find_index_eq (Bound (b - i)) args2) + 1) |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
845 |
(1 upto b) |
31514 | 846 |
val partial_mode = (1 upto k) \\ perm |
847 |
in |
|
848 |
if not (partial_mode subset is) then [] else |
|
849 |
let |
|
850 |
val is' = |
|
851 |
(fold_index (fn (i, j) => if j mem is then cons (i + 1) else I) perm []) |
|
852 |
|> fold (fn i => if i > k then cons (i - k + b) else I) is |
|
853 |
||
854 |
val res = map (fn x => Mode (m, is', x)) (cprods (map |
|
855 |
(fn (NONE, _) => [NONE] |
|
856 |
| (SOME js, arg) => map SOME (filter |
|
857 |
(fn Mode (_, js', _) => js=js') (modes_of_term modes arg))) |
|
858 |
(iss ~~ args1))) |
|
859 |
in res end |
|
860 |
end)) (AList.lookup op = modes name) |
|
861 |
in case strip_comb t' of |
|
862 |
(Const (name, _), args) => the_default default (mk_modes name args) |
|
863 |
| (Var ((name, _), _), args) => the (mk_modes name args) |
|
864 |
| (Free (name, _), args) => the (mk_modes name args) |
|
865 |
| _ => default end |
|
866 |
||
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
867 |
and |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
868 |
*) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
869 |
fun modes_of_term modes t = |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
870 |
let |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
871 |
val ks = 1 upto length (binder_types (fastype_of t)); |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
872 |
val default = [Mode (([], ks), ks, [])]; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
873 |
fun mk_modes name args = Option.map (maps (fn (m as (iss, is)) => |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
874 |
let |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
875 |
val (args1, args2) = |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
876 |
if length args < length iss then |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
877 |
error ("Too few arguments for inductive predicate " ^ name) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
878 |
else chop (length iss) args; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
879 |
val k = length args2; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
880 |
val prfx = 1 upto k |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
881 |
in |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
882 |
if not (is_prefix op = prfx is) then [] else |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
883 |
let val is' = map (fn i => i - k) (List.drop (is, k)) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
884 |
in map (fn x => Mode (m, is', x)) (cprods (map |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
885 |
(fn (NONE, _) => [NONE] |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
886 |
| (SOME js, arg) => map SOME (filter |
31170 | 887 |
(fn Mode (_, js', _) => js=js') (modes_of_term modes arg))) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
888 |
(iss ~~ args1))) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
889 |
end |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
890 |
end)) (AList.lookup op = modes name) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
891 |
|
32306
19f55947d4d5
removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents:
32287
diff
changeset
|
892 |
in |
19f55947d4d5
removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents:
32287
diff
changeset
|
893 |
case strip_comb (Envir.eta_contract t) of |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
894 |
(Const (name, _), args) => the_default default (mk_modes name args) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
895 |
| (Var ((name, _), _), args) => the (mk_modes name args) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
896 |
| (Free (name, _), args) => the (mk_modes name args) |
32306
19f55947d4d5
removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents:
32287
diff
changeset
|
897 |
| (Abs _, []) => error "Abs at param position" (* modes_of_param default modes t *) |
19f55947d4d5
removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents:
32287
diff
changeset
|
898 |
| _ => default |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
899 |
end |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
900 |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
901 |
fun select_mode_prem thy modes vs ps = |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
902 |
find_first (is_some o snd) (ps ~~ map |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
903 |
(fn Prem (us, t) => find_first (fn Mode (_, is, _) => |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
904 |
let |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
905 |
val (in_ts, out_ts) = split_smode is us; |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
906 |
val (out_ts', in_ts') = List.partition (is_constrt thy) out_ts; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
907 |
val vTs = maps term_vTs out_ts'; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
908 |
val dupTs = map snd (duplicates (op =) vTs) @ |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
909 |
List.mapPartial (AList.lookup (op =) vTs) vs; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
910 |
in |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
911 |
terms_vs (in_ts @ in_ts') subset vs andalso |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
912 |
forall (is_eqT o fastype_of) in_ts' andalso |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
913 |
term_vs t subset vs andalso |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
914 |
forall is_eqT dupTs |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
915 |
end) |
31170 | 916 |
(modes_of_term modes t handle Option => |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
917 |
error ("Bad predicate: " ^ Syntax.string_of_term_global thy t)) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
918 |
| Negprem (us, t) => find_first (fn Mode (_, is, _) => |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
919 |
length us = length is andalso |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
920 |
terms_vs us subset vs andalso |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
921 |
term_vs t subset vs) |
31170 | 922 |
(modes_of_term modes t handle Option => |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
923 |
error ("Bad predicate: " ^ Syntax.string_of_term_global thy t)) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
924 |
| Sidecond t => if term_vs t subset vs then SOME (Mode (([], []), [], [])) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
925 |
else NONE |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
926 |
) ps); |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
927 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
928 |
fun fold_prem f (Prem (args, _)) = fold f args |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
929 |
| fold_prem f (Negprem (args, _)) = fold f args |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
930 |
| fold_prem f (Sidecond t) = f t |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
931 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
932 |
fun all_subsets [] = [[]] |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
933 |
| all_subsets (x::xs) = let val xss' = all_subsets xs in xss' @ (map (cons x) xss') end |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
934 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
935 |
fun generator vTs v = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
936 |
let |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
937 |
val T = the (AList.lookup (op =) vTs v) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
938 |
in |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
939 |
(Generator (v, T), Mode (([], []), [], [])) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
940 |
end; |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
941 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
942 |
fun gen_prem (Prem (us, t)) = GeneratorPrem (us, t) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
943 |
| gen_prem _ = error "gen_prem : invalid input for gen_prem" |
32316
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
944 |
|
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
945 |
fun param_gen_prem param_vs (p as Prem (us, t as Free (v, _))) = |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
946 |
if member (op =) param_vs v then |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
947 |
GeneratorPrem (us, t) |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
948 |
else p |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
949 |
| param_gen_prem param_vs p = p |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
950 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
951 |
fun check_mode_clause with_generator thy param_vs modes gen_modes (iss, is) (ts, ps) = |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
952 |
let |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
953 |
val modes' = modes @ List.mapPartial |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
954 |
(fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)])) |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
955 |
(param_vs ~~ iss); |
32316
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
956 |
val gen_modes' = gen_modes @ List.mapPartial |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
957 |
(fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)])) |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
958 |
(param_vs ~~ iss); |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
959 |
val vTs = distinct (op =) ((fold o fold_prem) Term.add_frees ps (fold Term.add_frees ts [])) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
960 |
val prem_vs = distinct (op =) ((fold o fold_prem) Term.add_free_names ps []) |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
961 |
fun check_mode_prems acc_ps vs [] = SOME (acc_ps, vs) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
962 |
| check_mode_prems acc_ps vs ps = (case select_mode_prem thy modes' vs ps of |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
963 |
NONE => |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
964 |
(if with_generator then |
32316
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
965 |
(case select_mode_prem thy gen_modes' vs ps of |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
966 |
SOME (p, SOME mode) => check_mode_prems ((gen_prem p, mode) :: acc_ps) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
967 |
(case p of Prem (us, _) => vs union terms_vs us | _ => vs) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
968 |
(filter_out (equal p) ps) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
969 |
| NONE => |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
970 |
let |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
971 |
val all_generator_vs = all_subsets (prem_vs \\ vs) |> sort (int_ord o (pairself length)) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
972 |
in |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
973 |
case (find_first (fn generator_vs => is_some |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
974 |
(select_mode_prem thy modes' (vs union generator_vs) ps)) all_generator_vs) of |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
975 |
SOME generator_vs => check_mode_prems ((map (generator vTs) generator_vs) @ acc_ps) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
976 |
(vs union generator_vs) ps |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
977 |
| NONE => NONE |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
978 |
end) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
979 |
else |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
980 |
NONE) |
32316
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
981 |
| SOME (p, SOME mode) => check_mode_prems ((if with_generator then param_gen_prem param_vs p else p, mode) :: acc_ps) |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
982 |
(case p of Prem (us, _) => vs union terms_vs us | _ => vs) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
983 |
(filter_out (equal p) ps)) |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
984 |
val (in_ts, in_ts') = List.partition (is_constrt thy) (fst (split_smode is ts)); |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
985 |
val in_vs = terms_vs in_ts; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
986 |
val concl_vs = terms_vs ts |
32307
55166cd57a6d
exported functions for quickcheck generator; renamed type constructing functions to fit with the type name in the Predicate theory
bulwahn
parents:
32306
diff
changeset
|
987 |
in |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
988 |
if forall is_eqT (map snd (duplicates (op =) (maps term_vTs in_ts))) andalso |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
989 |
forall (is_eqT o fastype_of) in_ts' then |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
990 |
case check_mode_prems [] (param_vs union in_vs) ps of |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
991 |
NONE => NONE |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
992 |
| SOME (acc_ps, vs) => |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
993 |
if with_generator then |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
994 |
SOME (ts, (rev acc_ps) @ (map (generator vTs) (concl_vs \\ vs))) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
995 |
else |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
996 |
if concl_vs subset vs then SOME (ts, rev acc_ps) else NONE |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
997 |
else NONE |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
998 |
end; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
999 |
|
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1000 |
fun check_modes_pred with_generator thy param_vs preds modes gen_modes (p, ms) = |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1001 |
let val SOME rs = AList.lookup (op =) preds p |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1002 |
in (p, List.filter (fn m => case find_index |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1003 |
(is_none o check_mode_clause with_generator thy param_vs modes gen_modes m) rs of |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1004 |
~1 => true |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1005 |
| i => (Output.tracing ("Clause " ^ string_of_int (i + 1) ^ " of " ^ |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1006 |
p ^ " violates mode " ^ string_of_mode m); false)) ms) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1007 |
end; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1008 |
|
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1009 |
fun get_modes_pred with_generator thy param_vs preds modes gen_modes (p, ms) = |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1010 |
let |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1011 |
val SOME rs = AList.lookup (op =) preds p |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1012 |
in |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1013 |
(p, map (fn m => |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1014 |
(m, map (the o check_mode_clause with_generator thy param_vs modes gen_modes m) rs)) ms) |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1015 |
end; |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1016 |
|
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
1017 |
fun fixp f (x : (string * mode list) list) = |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1018 |
let val y = f x |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1019 |
in if x = y then x else fixp f y end; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1020 |
|
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1021 |
fun modes_of_arities arities = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1022 |
(map (fn (s, (ks, k)) => (s, cprod (cprods (map |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1023 |
(fn NONE => [NONE] |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1024 |
| SOME k' => map SOME (subsets 1 k')) ks), |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1025 |
subsets 1 k))) arities) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1026 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1027 |
fun infer_modes with_generator thy extra_modes arities param_vs preds = |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1028 |
let |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1029 |
val modes = |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1030 |
fixp (fn modes => |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1031 |
map (check_modes_pred with_generator thy param_vs preds (modes @ extra_modes) []) modes) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1032 |
(modes_of_arities arities) |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1033 |
in |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1034 |
map (get_modes_pred with_generator thy param_vs preds (modes @ extra_modes) []) modes |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1035 |
end; |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1036 |
|
32316
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
1037 |
fun remove_from rem [] = [] |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
1038 |
| remove_from rem ((k, vs) :: xs) = |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
1039 |
(case AList.lookup (op =) rem k of |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
1040 |
NONE => (k, vs) |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
1041 |
| SOME vs' => (k, vs \\ vs')) |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
1042 |
:: remove_from rem xs |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
1043 |
|
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1044 |
fun infer_modes_with_generator thy extra_modes arities param_vs preds = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1045 |
let |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1046 |
val prednames = map fst preds |
32316
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
1047 |
val extra_modes = all_modes_of thy |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1048 |
val gen_modes = all_generator_modes_of thy |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1049 |
|> filter_out (fn (name, _) => member (op =) prednames name) |
32316
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
1050 |
val starting_modes = remove_from extra_modes (modes_of_arities arities) |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1051 |
val modes = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1052 |
fixp (fn modes => |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1053 |
map (check_modes_pred true thy param_vs preds extra_modes (gen_modes @ modes)) modes) |
32316
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
1054 |
starting_modes |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1055 |
in |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1056 |
map (get_modes_pred true thy param_vs preds extra_modes (gen_modes @ modes)) modes |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1057 |
end; |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1058 |
|
31514 | 1059 |
(* term construction *) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1060 |
|
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1061 |
fun mk_v (names, vs) s T = (case AList.lookup (op =) vs s of |
31514 | 1062 |
NONE => (Free (s, T), (names, (s, [])::vs)) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1063 |
| SOME xs => |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1064 |
let |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1065 |
val s' = Name.variant names s; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1066 |
val v = Free (s', T) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1067 |
in |
31514 | 1068 |
(v, (s'::names, AList.update (op =) (s, v::xs) vs)) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1069 |
end); |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1070 |
|
31514 | 1071 |
fun distinct_v (Free (s, T)) nvs = mk_v nvs s T |
1072 |
| distinct_v (t $ u) nvs = |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1073 |
let |
31514 | 1074 |
val (t', nvs') = distinct_v t nvs; |
1075 |
val (u', nvs'') = distinct_v u nvs'; |
|
1076 |
in (t' $ u', nvs'') end |
|
1077 |
| distinct_v x nvs = (x, nvs); |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1078 |
|
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1079 |
fun compile_match thy compfuns eqs eqs' out_ts success_t = |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1080 |
let |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1081 |
val eqs'' = maps mk_eq eqs @ eqs' |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1082 |
val names = fold Term.add_free_names (success_t :: eqs'' @ out_ts) []; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1083 |
val name = Name.variant names "x"; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1084 |
val name' = Name.variant (name :: names) "y"; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1085 |
val T = mk_tupleT (map fastype_of out_ts); |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1086 |
val U = fastype_of success_t; |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1087 |
val U' = dest_predT compfuns U; |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1088 |
val v = Free (name, T); |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1089 |
val v' = Free (name', T); |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1090 |
in |
31723
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
31580
diff
changeset
|
1091 |
lambda v (fst (Datatype.make_case |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1092 |
(ProofContext.init thy) false [] v |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1093 |
[(mk_tuple out_ts, |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1094 |
if null eqs'' then success_t |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1095 |
else Const (@{const_name HOL.If}, HOLogic.boolT --> U --> U --> U) $ |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1096 |
foldr1 HOLogic.mk_conj eqs'' $ success_t $ |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1097 |
mk_bot compfuns U'), |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1098 |
(v', mk_bot compfuns U')])) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1099 |
end; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1100 |
|
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
1101 |
(*FIXME function can be removed*) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1102 |
fun mk_funcomp f t = |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1103 |
let |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1104 |
val names = Term.add_free_names t []; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1105 |
val Ts = binder_types (fastype_of t); |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1106 |
val vs = map Free |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1107 |
(Name.variant_list names (replicate (length Ts) "x") ~~ Ts) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1108 |
in |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1109 |
fold_rev lambda vs (f (list_comb (t, vs))) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1110 |
end; |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1111 |
(* |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1112 |
fun compile_param_ext thy compfuns modes (NONE, t) = t |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1113 |
| compile_param_ext thy compfuns modes (m as SOME (Mode ((iss, is'), is, ms)), t) = |
31514 | 1114 |
let |
1115 |
val (vs, u) = strip_abs t |
|
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1116 |
val (ivs, ovs) = split_mode is vs |
31514 | 1117 |
val (f, args) = strip_comb u |
1118 |
val (params, args') = chop (length ms) args |
|
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1119 |
val (inargs, outargs) = split_mode is' args' |
31514 | 1120 |
val b = length vs |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1121 |
val perm = map (fn i => (find_index_eq (Bound (b - i)) args') + 1) (1 upto b) |
31514 | 1122 |
val outp_perm = |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1123 |
snd (split_mode is perm) |
31514 | 1124 |
|> map (fn i => i - length (filter (fn x => x < i) is')) |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1125 |
val names = [] -- TODO |
31514 | 1126 |
val out_names = Name.variant_list names (replicate (length outargs) "x") |
1127 |
val f' = case f of |
|
1128 |
Const (name, T) => |
|
1129 |
if AList.defined op = modes name then |
|
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1130 |
mk_predfun_of thy compfuns (name, T) (iss, is') |
31514 | 1131 |
else error "compile param: Not an inductive predicate with correct mode" |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1132 |
| Free (name, T) => Free (name, param_funT_of compfuns T (SOME is')) |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1133 |
val outTs = dest_tupleT (dest_predT compfuns (body_type (fastype_of f'))) |
31514 | 1134 |
val out_vs = map Free (out_names ~~ outTs) |
1135 |
val params' = map (compile_param thy modes) (ms ~~ params) |
|
1136 |
val f_app = list_comb (f', params' @ inargs) |
|
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1137 |
val single_t = (mk_single compfuns (mk_tuple (map (fn i => nth out_vs (i - 1)) outp_perm))) |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1138 |
val match_t = compile_match thy compfuns [] [] out_vs single_t |
31514 | 1139 |
in list_abs (ivs, |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1140 |
mk_bind compfuns (f_app, match_t)) |
31514 | 1141 |
end |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1142 |
| compile_param_ext _ _ _ _ = error "compile params" |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1143 |
*) |
31514 | 1144 |
|
32315
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
1145 |
fun compile_param size thy compfuns (NONE, t) = t |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
1146 |
| compile_param size thy compfuns (m as SOME (Mode ((iss, is'), is, ms)), t) = |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1147 |
let |
31877
e3de75d3b898
exported is_registered; added debug messages; removed extended parameter compilation in predicate compiler
bulwahn
parents:
31876
diff
changeset
|
1148 |
val (f, args) = strip_comb (Envir.eta_contract t) |
31514 | 1149 |
val (params, args') = chop (length ms) args |
32315
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
1150 |
val params' = map (compile_param size thy compfuns) (ms ~~ params) |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
1151 |
val mk_fun_of = case size of NONE => mk_fun_of | SOME _ => mk_sizelim_fun_of |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
1152 |
val funT_of = case size of NONE => funT_of | SOME _ => sizelim_funT_of |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1153 |
val f' = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1154 |
case f of |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1155 |
Const (name, T) => |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1156 |
mk_fun_of compfuns thy (name, T) (iss, is') |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1157 |
| Free (name, T) => Free (name, funT_of compfuns (iss, is') T) |
32315
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
1158 |
| _ => error ("PredicateCompiler: illegal parameter term") |
31877
e3de75d3b898
exported is_registered; added debug messages; removed extended parameter compilation in predicate compiler
bulwahn
parents:
31876
diff
changeset
|
1159 |
in list_comb (f', params' @ args') end |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1160 |
|
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1161 |
fun compile_expr size thy ((Mode (mode, is, ms)), t) = |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1162 |
case strip_comb t of |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1163 |
(Const (name, T), params) => |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1164 |
let |
32315
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
1165 |
val params' = map (compile_param size thy PredicateCompFuns.compfuns) (ms ~~ params) |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
1166 |
val mk_fun_of = case size of NONE => mk_fun_of | SOME _ => mk_sizelim_fun_of |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1167 |
in |
32315
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
1168 |
list_comb (mk_fun_of PredicateCompFuns.compfuns thy (name, T) mode, params') |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1169 |
end |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1170 |
| (Free (name, T), args) => |
32315
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
1171 |
let |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
1172 |
val funT_of = case size of NONE => funT_of | SOME _ => sizelim_funT_of |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
1173 |
in |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
1174 |
list_comb (Free (name, funT_of PredicateCompFuns.compfuns ([], is) T), args) |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
1175 |
end; |
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
1176 |
|
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
1177 |
fun compile_gen_expr size thy compfuns ((Mode (mode, is, ms)), t) = |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1178 |
case strip_comb t of |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1179 |
(Const (name, T), params) => |
32316
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
1180 |
let |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
1181 |
val params' = map (compile_param size thy compfuns) (ms ~~ params) |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
1182 |
in |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
1183 |
list_comb (mk_generator_of compfuns thy (name, T) mode, params') |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
1184 |
end |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
1185 |
| (Free (name, T), args) => |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
1186 |
list_comb (Free (name, sizelim_funT_of RPredCompFuns.compfuns ([], is) T), args) |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
1187 |
|
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1188 |
(** specific rpred functions -- move them to the correct place in this file *) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1189 |
|
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1190 |
(* uncommented termify code; causes more trouble than expected at first *) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1191 |
(* |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1192 |
fun mk_valtermify_term (t as Const (c, T)) = HOLogic.mk_prod (t, Abs ("u", HOLogic.unitT, HOLogic.reflect_term t)) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1193 |
| mk_valtermify_term (Free (x, T)) = Free (x, termifyT T) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1194 |
| mk_valtermify_term (t1 $ t2) = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1195 |
let |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1196 |
val T = fastype_of t1 |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1197 |
val (T1, T2) = dest_funT T |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1198 |
val t1' = mk_valtermify_term t1 |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1199 |
val t2' = mk_valtermify_term t2 |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1200 |
in |
32657 | 1201 |
Const ("Code_Evaluation.valapp", termifyT T --> termifyT T1 --> termifyT T2) $ t1' $ t2' |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1202 |
end |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1203 |
| mk_valtermify_term _ = error "Not a valid term for mk_valtermify_term" |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1204 |
*) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1205 |
|
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1206 |
fun compile_clause compfuns size final_term thy all_vs param_vs (iss, is) inp (ts, moded_ps) = |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1207 |
let |
31514 | 1208 |
fun check_constrt t (names, eqs) = |
1209 |
if is_constrt thy t then (t, (names, eqs)) else |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1210 |
let |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1211 |
val s = Name.variant names "x"; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1212 |
val v = Free (s, fastype_of t) |
31514 | 1213 |
in (v, (s::names, HOLogic.mk_eq (v, t)::eqs)) end; |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1214 |
|
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1215 |
val (in_ts, out_ts) = split_smode is ts; |
31514 | 1216 |
val (in_ts', (all_vs', eqs)) = |
1217 |
fold_map check_constrt in_ts (all_vs, []); |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1218 |
|
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1219 |
fun compile_prems out_ts' vs names [] = |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1220 |
let |
31514 | 1221 |
val (out_ts'', (names', eqs')) = |
1222 |
fold_map check_constrt out_ts' (names, []); |
|
1223 |
val (out_ts''', (names'', constr_vs)) = fold_map distinct_v |
|
1224 |
out_ts'' (names', map (rpair []) vs); |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1225 |
in |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1226 |
(* termify code: |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1227 |
compile_match thy compfuns constr_vs (eqs @ eqs') out_ts''' |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1228 |
(mk_single compfuns (mk_tuple (map mk_valtermify_term out_ts))) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1229 |
*) |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1230 |
compile_match thy compfuns constr_vs (eqs @ eqs') out_ts''' |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1231 |
(final_term out_ts) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1232 |
end |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1233 |
| compile_prems out_ts vs names ((p, mode as Mode ((_, is), _, _)) :: ps) = |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1234 |
let |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1235 |
val vs' = distinct (op =) (flat (vs :: map term_vs out_ts)); |
31514 | 1236 |
val (out_ts', (names', eqs)) = |
1237 |
fold_map check_constrt out_ts (names, []) |
|
1238 |
val (out_ts'', (names'', constr_vs')) = fold_map distinct_v |
|
1239 |
out_ts' ((names', map (rpair []) vs)) |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1240 |
val (compiled_clause, rest) = case p of |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1241 |
Prem (us, t) => |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1242 |
let |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1243 |
val (in_ts, out_ts''') = split_smode is us; |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1244 |
val args = case size of |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1245 |
NONE => in_ts |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1246 |
| SOME size_t => in_ts @ [size_t] |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1247 |
val u = lift_pred compfuns |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1248 |
(list_comb (compile_expr size thy (mode, t), args)) |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1249 |
val rest = compile_prems out_ts''' vs' names'' ps |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1250 |
in |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1251 |
(u, rest) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1252 |
end |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1253 |
| Negprem (us, t) => |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1254 |
let |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1255 |
val (in_ts, out_ts''') = split_smode is us |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1256 |
val u = lift_pred compfuns |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1257 |
(mk_not PredicateCompFuns.compfuns (list_comb (compile_expr NONE thy (mode, t), in_ts))) |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1258 |
val rest = compile_prems out_ts''' vs' names'' ps |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1259 |
in |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1260 |
(u, rest) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1261 |
end |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1262 |
| Sidecond t => |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1263 |
let |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1264 |
val rest = compile_prems [] vs' names'' ps; |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1265 |
in |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1266 |
(mk_if compfuns t, rest) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1267 |
end |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1268 |
| GeneratorPrem (us, t) => |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1269 |
let |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1270 |
val (in_ts, out_ts''') = split_smode is us; |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1271 |
val args = case size of |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1272 |
NONE => in_ts |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1273 |
| SOME size_t => in_ts @ [size_t] |
32315
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
1274 |
val u = list_comb (compile_gen_expr size thy compfuns (mode, t), args) |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1275 |
val rest = compile_prems out_ts''' vs' names'' ps |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1276 |
in |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1277 |
(u, rest) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1278 |
end |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1279 |
| Generator (v, T) => |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1280 |
let |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1281 |
val u = lift_random (HOLogic.mk_random T @{term "1::code_numeral"}) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1282 |
val rest = compile_prems [Free (v, T)] vs' names'' ps; |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1283 |
in |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1284 |
(u, rest) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1285 |
end |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1286 |
in |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1287 |
compile_match thy compfuns constr_vs' eqs out_ts'' |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1288 |
(mk_bind compfuns (compiled_clause, rest)) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1289 |
end |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1290 |
val prem_t = compile_prems in_ts' param_vs all_vs' moded_ps; |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1291 |
in |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1292 |
mk_bind compfuns (mk_single compfuns inp, prem_t) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1293 |
end |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1294 |
|
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1295 |
fun compile_pred compfuns mk_fun_of use_size thy all_vs param_vs s T mode moded_cls = |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1296 |
let |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1297 |
val (Ts1, (Us1, Us2)) = split_mode mode (binder_types T) |
32315
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
1298 |
val funT_of = if use_size then sizelim_funT_of else funT_of |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1299 |
val Ts1' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) (fst mode) Ts1 |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1300 |
val xnames = Name.variant_list (all_vs @ param_vs) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1301 |
(map (fn i => "x" ^ string_of_int i) (snd mode)); |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1302 |
val size_name = Name.variant (all_vs @ param_vs @ xnames) "size" |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1303 |
(* termify code: val xs = map2 (fn s => fn T => Free (s, termifyT T)) xnames Us1; *) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1304 |
val xs = map2 (fn s => fn T => Free (s, T)) xnames Us1; |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1305 |
val params = map2 (fn s => fn T => Free (s, T)) param_vs Ts1' |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1306 |
val size = Free (size_name, @{typ "code_numeral"}) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1307 |
val decr_size = |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1308 |
if use_size then |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1309 |
SOME (Const ("HOL.minus_class.minus", @{typ "code_numeral => code_numeral => code_numeral"}) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1310 |
$ size $ Const ("HOL.one_class.one", @{typ "Code_Numeral.code_numeral"})) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1311 |
else |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1312 |
NONE |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1313 |
val cl_ts = |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1314 |
map (compile_clause compfuns decr_size (fn out_ts => mk_single compfuns (mk_tuple out_ts)) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1315 |
thy all_vs param_vs mode (mk_tuple xs)) moded_cls; |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1316 |
val t = foldr1 (mk_sup compfuns) cl_ts |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1317 |
val T' = mk_predT compfuns (mk_tupleT Us2) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1318 |
val size_t = Const (@{const_name "If"}, @{typ bool} --> T' --> T' --> T') |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1319 |
$ HOLogic.mk_eq (size, @{term "0 :: code_numeral"}) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1320 |
$ mk_bot compfuns (dest_predT compfuns T') $ t |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1321 |
val fun_const = mk_fun_of compfuns thy (s, T) mode |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1322 |
val eq = if use_size then |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1323 |
(list_comb (fun_const, params @ xs @ [size]), size_t) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1324 |
else |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1325 |
(list_comb (fun_const, params @ xs), t) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1326 |
in |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1327 |
HOLogic.mk_Trueprop (HOLogic.mk_eq eq) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1328 |
end; |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1329 |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1330 |
(* special setup for simpset *) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1331 |
val HOL_basic_ss' = HOL_basic_ss setSolver |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1332 |
(mk_solver "all_tac_solver" (fn _ => fn _ => all_tac)) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1333 |
|
31514 | 1334 |
(* Definition of executable functions and their intro and elim rules *) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1335 |
|
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30528
diff
changeset
|
1336 |
fun print_arities arities = tracing ("Arities:\n" ^ |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1337 |
cat_lines (map (fn (s, (ks, k)) => s ^ ": " ^ |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1338 |
space_implode " -> " (map |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1339 |
(fn NONE => "X" | SOME k' => string_of_int k') |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1340 |
(ks @ [SOME k]))) arities)); |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1341 |
|
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1342 |
fun mk_Eval_of ((x, T), NONE) names = (x, names) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1343 |
| mk_Eval_of ((x, T), SOME mode) names = let |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1344 |
val Ts = binder_types T |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1345 |
val argnames = Name.variant_list names |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1346 |
(map (fn i => "x" ^ string_of_int i) (1 upto (length Ts))); |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1347 |
val args = map Free (argnames ~~ Ts) |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1348 |
val (inargs, outargs) = split_smode mode args |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1349 |
val r = PredicateCompFuns.mk_Eval (list_comb (x, inargs), mk_tuple outargs) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1350 |
val t = fold_rev lambda args r |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1351 |
in |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1352 |
(t, argnames @ names) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1353 |
end; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1354 |
|
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1355 |
fun create_intro_elim_rule (mode as (iss, is)) defthm mode_id funT pred thy = |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1356 |
let |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1357 |
val Ts = binder_types (fastype_of pred) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1358 |
val funtrm = Const (mode_id, funT) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1359 |
val argnames = Name.variant_list [] |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1360 |
(map (fn i => "x" ^ string_of_int i) (1 upto (length Ts))); |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1361 |
val (Ts1, Ts2) = chop (length iss) Ts; |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1362 |
val Ts1' = map2 (fn NONE => I | SOME is => funT_of (PredicateCompFuns.compfuns) ([], is)) iss Ts1 |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1363 |
val args = map Free (argnames ~~ (Ts1' @ Ts2)) |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1364 |
val (params, ioargs) = chop (length iss) args |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1365 |
val (inargs, outargs) = split_smode is ioargs |
31550
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31549
diff
changeset
|
1366 |
val param_names = Name.variant_list argnames |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1367 |
(map (fn i => "p" ^ string_of_int i) (1 upto (length iss))) |
31550
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31549
diff
changeset
|
1368 |
val param_vs = map Free (param_names ~~ Ts1) |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1369 |
val (params', names) = fold_map mk_Eval_of ((params ~~ Ts1) ~~ iss) [] |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1370 |
val predpropI = HOLogic.mk_Trueprop (list_comb (pred, param_vs @ ioargs)) |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1371 |
val predpropE = HOLogic.mk_Trueprop (list_comb (pred, params' @ ioargs)) |
31550
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31549
diff
changeset
|
1372 |
val param_eqs = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (param_vs ~~ params') |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1373 |
val funargs = params @ inargs |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1374 |
val funpropE = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, funargs), |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1375 |
if null outargs then Free("y", HOLogic.unitT) else mk_tuple outargs)) |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1376 |
val funpropI = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, funargs), |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1377 |
mk_tuple outargs)) |
31550
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31549
diff
changeset
|
1378 |
val introtrm = Logic.list_implies (predpropI :: param_eqs, funpropI) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1379 |
val simprules = [defthm, @{thm eval_pred}, |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1380 |
@{thm "split_beta"}, @{thm "fst_conv"}, @{thm "snd_conv"}] |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1381 |
val unfolddef_tac = Simplifier.asm_full_simp_tac (HOL_basic_ss addsimps simprules) 1 |
31550
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31549
diff
changeset
|
1382 |
val introthm = Goal.prove (ProofContext.init thy) (argnames @ param_names @ ["y"]) [] introtrm (fn {...} => unfolddef_tac) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1383 |
val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT)); |
31550
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31549
diff
changeset
|
1384 |
val elimtrm = Logic.list_implies ([funpropE, Logic.mk_implies (predpropE, P)], P) |
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31549
diff
changeset
|
1385 |
val elimthm = Goal.prove (ProofContext.init thy) (argnames @ param_names @ ["y", "P"]) [] elimtrm (fn {...} => unfolddef_tac) |
31514 | 1386 |
in |
1387 |
(introthm, elimthm) |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1388 |
end; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1389 |
|
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1390 |
fun create_constname_of_mode thy prefix name mode = |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1391 |
let |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1392 |
fun string_of_mode mode = if null mode then "0" |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1393 |
else space_implode "_" (map string_of_int mode) |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1394 |
val HOmode = space_implode "_and_" |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1395 |
(fold (fn NONE => I | SOME mode => cons (string_of_mode mode)) (fst mode) []) |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1396 |
in |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1397 |
(Sign.full_bname thy (prefix ^ (Long_Name.base_name name))) ^ |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1398 |
(if HOmode = "" then "_" else "_for_" ^ HOmode ^ "_yields_") ^ (string_of_mode (snd mode)) |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1399 |
end; |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1400 |
|
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1401 |
fun create_definitions preds (name, modes) thy = |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1402 |
let |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1403 |
val compfuns = PredicateCompFuns.compfuns |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1404 |
val T = AList.lookup (op =) preds name |> the |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1405 |
fun create_definition (mode as (iss, is)) thy = let |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1406 |
val mode_cname = create_constname_of_mode thy "" name mode |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1407 |
val mode_cbasename = Long_Name.base_name mode_cname |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1408 |
val Ts = binder_types T |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1409 |
val (Ts1, Ts2) = chop (length iss) Ts |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1410 |
val (Us1, Us2) = split_smode is Ts2 |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1411 |
val Ts1' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) iss Ts1 |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1412 |
val funT = (Ts1' @ Us1) ---> (mk_predT compfuns (mk_tupleT Us2)) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1413 |
val names = Name.variant_list [] |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1414 |
(map (fn i => "x" ^ string_of_int i) (1 upto (length Ts))); |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1415 |
val xs = map Free (names ~~ (Ts1' @ Ts2)); |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1416 |
val (xparams, xargs) = chop (length iss) xs; |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1417 |
val (xins, xouts) = split_smode is xargs |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1418 |
val (xparams', names') = fold_map mk_Eval_of ((xparams ~~ Ts1) ~~ iss) names |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1419 |
fun mk_split_lambda [] t = lambda (Free (Name.variant names' "x", HOLogic.unitT)) t |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1420 |
| mk_split_lambda [x] t = lambda x t |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1421 |
| mk_split_lambda xs t = |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1422 |
let |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1423 |
fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t)) |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1424 |
| mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t)) |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1425 |
in |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1426 |
mk_split_lambda' xs t |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1427 |
end; |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1428 |
val predterm = PredicateCompFuns.mk_Enum (mk_split_lambda xouts |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1429 |
(list_comb (Const (name, T), xparams' @ xargs))) |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1430 |
val lhs = list_comb (Const (mode_cname, funT), xparams @ xins) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1431 |
val def = Logic.mk_equals (lhs, predterm) |
31514 | 1432 |
val ([definition], thy') = thy |> |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1433 |
Sign.add_consts_i [(Binding.name mode_cbasename, funT, NoSyn)] |> |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1434 |
PureThy.add_defs false [((Binding.name (mode_cbasename ^ "_def"), def), [])] |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1435 |
val (intro, elim) = |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1436 |
create_intro_elim_rule mode definition mode_cname funT (Const (name, T)) thy' |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1437 |
in thy' |> add_predfun name mode (mode_cname, definition, intro, elim) |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1438 |
|> PureThy.store_thm (Binding.name (mode_cbasename ^ "I"), intro) |> snd |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1439 |
|> PureThy.store_thm (Binding.name (mode_cbasename ^ "E"), elim) |> snd |
31554
a6c6bf2751a0
added sporadic (Local)Theory.checkpoint, to enable parallel proof checking;
wenzelm
parents:
31550
diff
changeset
|
1440 |
|> Theory.checkpoint |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1441 |
end; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1442 |
in |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1443 |
fold create_definition modes thy |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1444 |
end; |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1445 |
|
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1446 |
fun sizelim_create_definitions preds (name, modes) thy = |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1447 |
let |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1448 |
val T = AList.lookup (op =) preds name |> the |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1449 |
fun create_definition mode thy = |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1450 |
let |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1451 |
val mode_cname = create_constname_of_mode thy "sizelim_" name mode |
32315
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
1452 |
val funT = sizelim_funT_of PredicateCompFuns.compfuns mode T |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1453 |
in |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1454 |
thy |> Sign.add_consts_i [(Binding.name (Long_Name.base_name mode_cname), funT, NoSyn)] |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1455 |
|> set_sizelim_function_name name mode mode_cname |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1456 |
end; |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1457 |
in |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1458 |
fold create_definition modes thy |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1459 |
end; |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1460 |
|
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1461 |
fun rpred_create_definitions preds (name, modes) thy = |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1462 |
let |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1463 |
val T = AList.lookup (op =) preds name |> the |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1464 |
fun create_definition mode thy = |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1465 |
let |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1466 |
val mode_cname = create_constname_of_mode thy "gen_" name mode |
32316
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
1467 |
val funT = sizelim_funT_of RPredCompFuns.compfuns mode T |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1468 |
in |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1469 |
thy |> Sign.add_consts_i [(Binding.name (Long_Name.base_name mode_cname), funT, NoSyn)] |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1470 |
|> set_generator_name name mode mode_cname |
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1471 |
end; |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1472 |
in |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1473 |
fold create_definition modes thy |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1474 |
end; |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1475 |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1476 |
(* Proving equivalence of term *) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1477 |
|
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1478 |
fun is_Type (Type _) = true |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1479 |
| is_Type _ = false |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1480 |
|
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1481 |
(* returns true if t is an application of an datatype constructor *) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1482 |
(* which then consequently would be splitted *) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1483 |
(* else false *) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1484 |
fun is_constructor thy t = |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1485 |
if (is_Type (fastype_of t)) then |
31784 | 1486 |
(case Datatype.get_info thy ((fst o dest_Type o fastype_of) t) of |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1487 |
NONE => false |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1488 |
| SOME info => (let |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1489 |
val constr_consts = maps (fn (_, (_, _, constrs)) => map fst constrs) (#descr info) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1490 |
val (c, _) = strip_comb t |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1491 |
in (case c of |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1492 |
Const (name, _) => name mem_string constr_consts |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1493 |
| _ => false) end)) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1494 |
else false |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1495 |
|
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1496 |
(* MAJOR FIXME: prove_params should be simple |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1497 |
- different form of introrule for parameters ? *) |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1498 |
fun prove_param thy (NONE, t) = TRY (rtac @{thm refl} 1) |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1499 |
| prove_param thy (m as SOME (Mode (mode, is, ms)), t) = |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1500 |
let |
32315
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
1501 |
val (f, args) = strip_comb (Envir.eta_contract t) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1502 |
val (params, _) = chop (length ms) args |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1503 |
val f_tac = case f of |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1504 |
Const (name, T) => simp_tac (HOL_basic_ss addsimps |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1505 |
(@{thm eval_pred}::(predfun_definition_of thy name mode):: |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1506 |
@{thm "Product_Type.split_conv"}::[])) 1 |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1507 |
| Free _ => TRY (rtac @{thm refl} 1) |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1508 |
| Abs _ => error "prove_param: No valid parameter term" |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1509 |
in |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1510 |
REPEAT_DETERM (etac @{thm thin_rl} 1) |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1511 |
THEN REPEAT_DETERM (rtac @{thm ext} 1) |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1512 |
THEN print_tac "prove_param" |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1513 |
THEN f_tac |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1514 |
THEN print_tac "after simplification in prove_args" |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1515 |
THEN (EVERY (map (prove_param thy) (ms ~~ params))) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1516 |
THEN (REPEAT_DETERM (atac 1)) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1517 |
end |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1518 |
|
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1519 |
fun prove_expr thy (Mode (mode, is, ms), t, us) (premposition : int) = |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1520 |
case strip_comb t of |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1521 |
(Const (name, T), args) => |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1522 |
let |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1523 |
val introrule = predfun_intro_of thy name mode |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1524 |
val (args1, args2) = chop (length ms) args |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1525 |
in |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1526 |
rtac @{thm bindI} 1 |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1527 |
THEN print_tac "before intro rule:" |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1528 |
(* for the right assumption in first position *) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1529 |
THEN rotate_tac premposition 1 |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1530 |
THEN debug_tac (Display.string_of_thm (ProofContext.init thy) introrule) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1531 |
THEN rtac introrule 1 |
32306
19f55947d4d5
removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents:
32287
diff
changeset
|
1532 |
THEN print_tac "after intro rule" |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1533 |
(* work with parameter arguments *) |
31550
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31549
diff
changeset
|
1534 |
THEN (atac 1) |
32306
19f55947d4d5
removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents:
32287
diff
changeset
|
1535 |
THEN (print_tac "parameter goal") |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1536 |
THEN (EVERY (map (prove_param thy) (ms ~~ args1))) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1537 |
THEN (REPEAT_DETERM (atac 1)) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1538 |
end |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1539 |
| _ => rtac @{thm bindI} 1 THEN atac 1 |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1540 |
|
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1541 |
fun SOLVED tac st = FILTER (fn st' => nprems_of st' = nprems_of st - 1) tac st; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1542 |
|
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1543 |
fun SOLVEDALL tac st = FILTER (fn st' => nprems_of st' = 0) tac st |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1544 |
|
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1545 |
fun prove_match thy (out_ts : term list) = let |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1546 |
fun get_case_rewrite t = |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1547 |
if (is_constructor thy t) then let |
31784 | 1548 |
val case_rewrites = (#case_rewrites (Datatype.the_info thy |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1549 |
((fst o dest_Type o fastype_of) t))) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1550 |
in case_rewrites @ (flat (map get_case_rewrite (snd (strip_comb t)))) end |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1551 |
else [] |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1552 |
val simprules = @{thm "unit.cases"} :: @{thm "prod.cases"} :: (flat (map get_case_rewrite out_ts)) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1553 |
(* replace TRY by determining if it necessary - are there equations when calling compile match? *) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1554 |
in |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1555 |
(* make this simpset better! *) |
31515 | 1556 |
asm_simp_tac (HOL_basic_ss' addsimps simprules) 1 |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1557 |
THEN print_tac "after prove_match:" |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1558 |
THEN (DETERM (TRY (EqSubst.eqsubst_tac (ProofContext.init thy) [0] [@{thm "HOL.if_P"}] 1 |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1559 |
THEN (REPEAT_DETERM (rtac @{thm conjI} 1 THEN (SOLVED (asm_simp_tac HOL_basic_ss 1)))) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1560 |
THEN (SOLVED (asm_simp_tac HOL_basic_ss 1))))) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1561 |
THEN print_tac "after if simplification" |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1562 |
end; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1563 |
|
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1564 |
(* corresponds to compile_fun -- maybe call that also compile_sidecond? *) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1565 |
|
31515 | 1566 |
fun prove_sidecond thy modes t = |
1567 |
let |
|
1568 |
fun preds_of t nameTs = case strip_comb t of |
|
1569 |
(f as Const (name, T), args) => |
|
1570 |
if AList.defined (op =) modes name then (name, T) :: nameTs |
|
1571 |
else fold preds_of args nameTs |
|
1572 |
| _ => nameTs |
|
1573 |
val preds = preds_of t [] |
|
1574 |
val defs = map |
|
1575 |
(fn (pred, T) => predfun_definition_of thy pred ([], (1 upto (length (binder_types T))))) |
|
1576 |
preds |
|
1577 |
in |
|
1578 |
(* remove not_False_eq_True when simpset in prove_match is better *) |
|
1579 |
simp_tac (HOL_basic_ss addsimps @{thm not_False_eq_True} :: @{thm eval_pred} :: defs) 1 |
|
1580 |
(* need better control here! *) |
|
1581 |
end |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1582 |
|
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1583 |
fun prove_clause thy nargs modes (iss, is) (_, clauses) (ts, moded_ps) = |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1584 |
let |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1585 |
val (in_ts, clause_out_ts) = split_smode is ts; |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1586 |
fun prove_prems out_ts [] = |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1587 |
(prove_match thy out_ts) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1588 |
THEN asm_simp_tac HOL_basic_ss' 1 |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1589 |
THEN (rtac (if null clause_out_ts then @{thm singleI_unit} else @{thm singleI}) 1) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1590 |
| prove_prems out_ts ((p, mode as Mode ((iss, is), _, param_modes)) :: ps) = |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1591 |
let |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1592 |
val premposition = (find_index (equal p) clauses) + nargs |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1593 |
val rest_tac = (case p of Prem (us, t) => |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1594 |
let |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1595 |
val (_, out_ts''') = split_smode is us |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1596 |
val rec_tac = prove_prems out_ts''' ps |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1597 |
in |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1598 |
print_tac "before clause:" |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1599 |
THEN asm_simp_tac HOL_basic_ss 1 |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1600 |
THEN print_tac "before prove_expr:" |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1601 |
THEN prove_expr thy (mode, t, us) premposition |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1602 |
THEN print_tac "after prove_expr:" |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1603 |
THEN rec_tac |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1604 |
end |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1605 |
| Negprem (us, t) => |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1606 |
let |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1607 |
val (_, out_ts''') = split_smode is us |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1608 |
val rec_tac = prove_prems out_ts''' ps |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1609 |
val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1610 |
val (_, params) = strip_comb t |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1611 |
in |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1612 |
rtac @{thm bindI} 1 |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1613 |
THEN (if (is_some name) then |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1614 |
simp_tac (HOL_basic_ss addsimps [predfun_definition_of thy (the name) (iss, is)]) 1 |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1615 |
THEN rtac @{thm not_predI} 1 |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1616 |
THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1 |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1617 |
THEN (REPEAT_DETERM (atac 1)) |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1618 |
(* FIXME: work with parameter arguments *) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1619 |
THEN (EVERY (map (prove_param thy) (param_modes ~~ params))) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1620 |
else |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1621 |
rtac @{thm not_predI'} 1) |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1622 |
THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1 |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1623 |
THEN rec_tac |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1624 |
end |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1625 |
| Sidecond t => |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1626 |
rtac @{thm bindI} 1 |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1627 |
THEN rtac @{thm if_predI} 1 |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1628 |
THEN print_tac "before sidecond:" |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1629 |
THEN prove_sidecond thy modes t |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1630 |
THEN print_tac "after sidecond:" |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1631 |
THEN prove_prems [] ps) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1632 |
in (prove_match thy out_ts) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1633 |
THEN rest_tac |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1634 |
end; |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1635 |
val prems_tac = prove_prems in_ts moded_ps |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1636 |
in |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1637 |
rtac @{thm bindI} 1 |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1638 |
THEN rtac @{thm singleI} 1 |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1639 |
THEN prems_tac |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1640 |
end; |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1641 |
|
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1642 |
fun select_sup 1 1 = [] |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1643 |
| select_sup _ 1 = [rtac @{thm supI1}] |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1644 |
| select_sup n i = (rtac @{thm supI2})::(select_sup (n - 1) (i - 1)); |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1645 |
|
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1646 |
fun prove_one_direction thy clauses preds modes pred mode moded_clauses = |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1647 |
let |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1648 |
val T = the (AList.lookup (op =) preds pred) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1649 |
val nargs = length (binder_types T) - nparams_of thy pred |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1650 |
val pred_case_rule = the_elim_of thy pred |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1651 |
in |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1652 |
REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"})) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1653 |
THEN etac (predfun_elim_of thy pred mode) 1 |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1654 |
THEN etac pred_case_rule 1 |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1655 |
THEN (EVERY (map |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1656 |
(fn i => EVERY' (select_sup (length moded_clauses) i) i) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1657 |
(1 upto (length moded_clauses)))) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1658 |
THEN (EVERY (map2 (prove_clause thy nargs modes mode) clauses moded_clauses)) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1659 |
THEN print_tac "proved one direction" |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1660 |
end; |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1661 |
|
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1662 |
(** Proof in the other direction **) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1663 |
|
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1664 |
fun prove_match2 thy out_ts = let |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1665 |
fun split_term_tac (Free _) = all_tac |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1666 |
| split_term_tac t = |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1667 |
if (is_constructor thy t) then let |
31784 | 1668 |
val info = Datatype.the_info thy ((fst o dest_Type o fastype_of) t) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1669 |
val num_of_constrs = length (#case_rewrites info) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1670 |
(* special treatment of pairs -- because of fishing *) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1671 |
val split_rules = case (fst o dest_Type o fastype_of) t of |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1672 |
"*" => [@{thm prod.split_asm}] |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1673 |
| _ => PureThy.get_thms thy (((fst o dest_Type o fastype_of) t) ^ ".split_asm") |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1674 |
val (_, ts) = strip_comb t |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1675 |
in |
31515 | 1676 |
(Splitter.split_asm_tac split_rules 1) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1677 |
(* THEN (Simplifier.asm_full_simp_tac HOL_basic_ss 1) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1678 |
THEN (DETERM (TRY (etac @{thm Pair_inject} 1))) *) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1679 |
THEN (REPEAT_DETERM_N (num_of_constrs - 1) (etac @{thm botE} 1 ORELSE etac @{thm botE} 2)) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1680 |
THEN (EVERY (map split_term_tac ts)) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1681 |
end |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1682 |
else all_tac |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1683 |
in |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1684 |
split_term_tac (mk_tuple out_ts) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1685 |
THEN (DETERM (TRY ((Splitter.split_asm_tac [@{thm "split_if_asm"}] 1) THEN (etac @{thm botE} 2)))) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1686 |
end |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1687 |
|
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1688 |
(* VERY LARGE SIMILIRATIY to function prove_param |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1689 |
-- join both functions |
31550
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31549
diff
changeset
|
1690 |
*) |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1691 |
(* TODO: remove function *) |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1692 |
|
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1693 |
fun prove_param2 thy (NONE, t) = all_tac |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1694 |
| prove_param2 thy (m as SOME (Mode (mode, is, ms)), t) = let |
32315
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
1695 |
val (f, args) = strip_comb (Envir.eta_contract t) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1696 |
val (params, _) = chop (length ms) args |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1697 |
val f_tac = case f of |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1698 |
Const (name, T) => full_simp_tac (HOL_basic_ss addsimps |
31514 | 1699 |
(@{thm eval_pred}::(predfun_definition_of thy name mode) |
1700 |
:: @{thm "Product_Type.split_conv"}::[])) 1 |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1701 |
| Free _ => all_tac |
32315
79f324944be4
tuned proof procedure; added size-limiting predicate compilation of higher order predicates; added guessing of number parameters for registrating predicates; removed debug messages
bulwahn
parents:
32314
diff
changeset
|
1702 |
| _ => error "prove_param2: illegal parameter term" |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1703 |
in |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1704 |
print_tac "before simplification in prove_args:" |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1705 |
THEN f_tac |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1706 |
THEN print_tac "after simplification in prove_args" |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1707 |
THEN (EVERY (map (prove_param2 thy) (ms ~~ params))) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1708 |
end |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1709 |
|
31877
e3de75d3b898
exported is_registered; added debug messages; removed extended parameter compilation in predicate compiler
bulwahn
parents:
31876
diff
changeset
|
1710 |
|
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1711 |
fun prove_expr2 thy (Mode (mode, is, ms), t) = |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1712 |
(case strip_comb t of |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1713 |
(Const (name, T), args) => |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1714 |
etac @{thm bindE} 1 |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1715 |
THEN (REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"}))) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1716 |
THEN print_tac "prove_expr2-before" |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1717 |
THEN (debug_tac (Syntax.string_of_term_global thy |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1718 |
(prop_of (predfun_elim_of thy name mode)))) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1719 |
THEN (etac (predfun_elim_of thy name mode) 1) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1720 |
THEN print_tac "prove_expr2" |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1721 |
THEN (EVERY (map (prove_param2 thy) (ms ~~ args))) |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1722 |
THEN print_tac "finished prove_expr2" |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1723 |
| _ => etac @{thm bindE} 1) |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1724 |
|
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1725 |
(* FIXME: what is this for? *) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1726 |
(* replace defined by has_mode thy pred *) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1727 |
(* TODO: rewrite function *) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1728 |
fun prove_sidecond2 thy modes t = let |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1729 |
fun preds_of t nameTs = case strip_comb t of |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1730 |
(f as Const (name, T), args) => |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1731 |
if AList.defined (op =) modes name then (name, T) :: nameTs |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1732 |
else fold preds_of args nameTs |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1733 |
| _ => nameTs |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1734 |
val preds = preds_of t [] |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1735 |
val defs = map |
31514 | 1736 |
(fn (pred, T) => predfun_definition_of thy pred ([], (1 upto (length (binder_types T))))) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1737 |
preds |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1738 |
in |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1739 |
(* only simplify the one assumption *) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1740 |
full_simp_tac (HOL_basic_ss' addsimps @{thm eval_pred} :: defs) 1 |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1741 |
(* need better control here! *) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1742 |
THEN print_tac "after sidecond2 simplification" |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1743 |
end |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1744 |
|
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1745 |
fun prove_clause2 thy modes pred (iss, is) (ts, ps) i = |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1746 |
let |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1747 |
val pred_intro_rule = nth (intros_of thy pred) (i - 1) |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1748 |
val (in_ts, clause_out_ts) = split_smode is ts; |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1749 |
fun prove_prems2 out_ts [] = |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1750 |
print_tac "before prove_match2 - last call:" |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1751 |
THEN prove_match2 thy out_ts |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1752 |
THEN print_tac "after prove_match2 - last call:" |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1753 |
THEN (etac @{thm singleE} 1) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1754 |
THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1)) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1755 |
THEN (asm_full_simp_tac HOL_basic_ss' 1) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1756 |
THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1)) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1757 |
THEN (asm_full_simp_tac HOL_basic_ss' 1) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1758 |
THEN SOLVED (print_tac "state before applying intro rule:" |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1759 |
THEN (rtac pred_intro_rule 1) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1760 |
(* How to handle equality correctly? *) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1761 |
THEN (print_tac "state before assumption matching") |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1762 |
THEN (REPEAT (atac 1 ORELSE |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1763 |
(CHANGED (asm_full_simp_tac HOL_basic_ss' 1) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1764 |
THEN print_tac "state after simp_tac:")))) |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1765 |
| prove_prems2 out_ts ((p, mode as Mode ((iss, is), _, param_modes)) :: ps) = |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1766 |
let |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1767 |
val rest_tac = (case p of |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1768 |
Prem (us, t) => |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1769 |
let |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1770 |
val (_, out_ts''') = split_smode is us |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1771 |
val rec_tac = prove_prems2 out_ts''' ps |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1772 |
in |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1773 |
(prove_expr2 thy (mode, t)) THEN rec_tac |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1774 |
end |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1775 |
| Negprem (us, t) => |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1776 |
let |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1777 |
val (_, out_ts''') = split_smode is us |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1778 |
val rec_tac = prove_prems2 out_ts''' ps |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1779 |
val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1780 |
val (_, params) = strip_comb t |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1781 |
in |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1782 |
print_tac "before neg prem 2" |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1783 |
THEN etac @{thm bindE} 1 |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1784 |
THEN (if is_some name then |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1785 |
full_simp_tac (HOL_basic_ss addsimps [predfun_definition_of thy (the name) (iss, is)]) 1 |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1786 |
THEN etac @{thm not_predE} 1 |
32313
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1787 |
THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1 |
a984c04927b4
refactoring predicate compiler; repaired proof procedure to handle all test cases
bulwahn
parents:
32312
diff
changeset
|
1788 |
THEN (EVERY (map (prove_param2 thy) (param_modes ~~ params))) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1789 |
else |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1790 |
etac @{thm not_predE'} 1) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1791 |
THEN rec_tac |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1792 |
end |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1793 |
| Sidecond t => |
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1794 |
etac @{thm bindE} 1 |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1795 |
THEN etac @{thm if_predE} 1 |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1796 |
THEN prove_sidecond2 thy modes t |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1797 |
THEN prove_prems2 [] ps) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1798 |
in print_tac "before prove_match2:" |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1799 |
THEN prove_match2 thy out_ts |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1800 |
THEN print_tac "after prove_match2:" |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1801 |
THEN rest_tac |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1802 |
end; |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1803 |
val prems_tac = prove_prems2 in_ts ps |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1804 |
in |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1805 |
print_tac "starting prove_clause2" |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1806 |
THEN etac @{thm bindE} 1 |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1807 |
THEN (etac @{thm singleE'} 1) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1808 |
THEN (TRY (etac @{thm Pair_inject} 1)) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1809 |
THEN print_tac "after singleE':" |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1810 |
THEN prems_tac |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1811 |
end; |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1812 |
|
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1813 |
fun prove_other_direction thy modes pred mode moded_clauses = |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1814 |
let |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1815 |
fun prove_clause clause i = |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1816 |
(if i < length moded_clauses then etac @{thm supE} 1 else all_tac) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1817 |
THEN (prove_clause2 thy modes pred mode clause i) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1818 |
in |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1819 |
(DETERM (TRY (rtac @{thm unit.induct} 1))) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1820 |
THEN (REPEAT_DETERM (CHANGED (rewtac @{thm split_paired_all}))) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1821 |
THEN (rtac (predfun_intro_of thy pred mode) 1) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1822 |
THEN (REPEAT_DETERM (rtac @{thm refl} 2)) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1823 |
THEN (EVERY (map2 prove_clause moded_clauses (1 upto (length moded_clauses)))) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1824 |
end; |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1825 |
|
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1826 |
(** proof procedure **) |
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1827 |
|
32309
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1828 |
fun prove_pred thy clauses preds modes pred mode (moded_clauses, compiled_term) = |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1829 |
let |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1830 |
val ctxt = ProofContext.init thy |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1831 |
val clauses = the (AList.lookup (op =) clauses pred) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1832 |
in |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1833 |
Goal.prove ctxt (Term.add_free_names compiled_term []) [] compiled_term |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1834 |
(if !do_proofs then |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1835 |
(fn _ => |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1836 |
rtac @{thm pred_iffI} 1 |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1837 |
THEN prove_one_direction thy clauses preds modes pred mode moded_clauses |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1838 |
THEN print_tac "proved one direction" |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1839 |
THEN prove_other_direction thy modes pred mode moded_clauses |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1840 |
THEN print_tac "proved other direction") |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1841 |
else (fn _ => mycheat_tac thy 1)) |
4a18f3cf6362
imported patch changed mode inference of predicate compiler to return infered dataflow
bulwahn
parents:
32308
diff
changeset
|
1842 |
end; |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1843 |
|
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1844 |
(* composition of mode inference, definition, compilation and proof *) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1845 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1846 |
(** auxillary combinators for table of preds and modes **) |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1847 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1848 |
fun map_preds_modes f preds_modes_table = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1849 |
map (fn (pred, modes) => |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1850 |
(pred, map (fn (mode, value) => (mode, f pred mode value)) modes)) preds_modes_table |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1851 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1852 |
fun join_preds_modes table1 table2 = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1853 |
map_preds_modes (fn pred => fn mode => fn value => |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1854 |
(value, the (AList.lookup (op =) (the (AList.lookup (op =) table2 pred)) mode))) table1 |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1855 |
|
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1856 |
fun maps_modes preds_modes_table = |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1857 |
map (fn (pred, modes) => |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1858 |
(pred, map (fn (mode, value) => value) modes)) preds_modes_table |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1859 |
|
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1860 |
fun compile_preds compfuns mk_fun_of use_size thy all_vs param_vs preds moded_clauses = |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1861 |
map_preds_modes (fn pred => compile_pred compfuns mk_fun_of use_size thy all_vs param_vs pred |
32310
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1862 |
(the (AList.lookup (op =) preds pred))) moded_clauses |
89f3c616a2d1
extended mode inference by adding generators; adopted compilation; added some functions for constructing term closures
bulwahn
parents:
32309
diff
changeset
|
1863 |
|
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1864 |
fun prove thy clauses preds modes moded_clauses compiled_terms = |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1865 |
map_preds_modes (prove_pred thy clauses preds modes) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1866 |
(join_preds_modes moded_clauses compiled_terms) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1867 |
|
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1868 |
fun prove_by_skip thy _ _ _ _ compiled_terms = |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1869 |
map_preds_modes (fn pred => fn mode => fn t => Drule.standard (SkipProof.make_thm thy t)) |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1870 |
compiled_terms |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1871 |
|
31514 | 1872 |
fun prepare_intrs thy prednames = |
1873 |
let |
|
32308
c2b74affab85
imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
32307
diff
changeset
|
1874 |
val intrs = maps (intros_of thy) prednames |
31514 | 1875 |
|> map (Logic.unvarify o prop_of) |
1876 |
val nparams = nparams_of thy (hd prednames) |
|
32316
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
1877 |
val extra_modes = all_modes_of thy |> filter_out (fn (name, _) => member (op =) prednames name) |
31514 | 1878 |
val preds = distinct (op =) (map (dest_Const o fst o (strip_intro_concl nparams)) intrs) |
1879 |
val _ $ u = Logic.strip_imp_concl (hd intrs); |
|
1880 |
val params = List.take (snd (strip_comb u), nparams); |
|
1881 |
val param_vs = maps term_vs params |
|
1882 |
val all_vs = terms_vs intrs |
|
1883 |
fun dest_prem t = |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1884 |
(case strip_comb t of |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1885 |
(v as Free _, ts) => if v mem params then Prem (ts, v) else Sidecond t |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1886 |
| (c as Const (@{const_name Not}, _), [t]) => (case dest_prem t of |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1887 |
Prem (ts, t) => Negprem (ts, t) |
31515 | 1888 |
| Negprem _ => error ("Double negation not allowed in premise: " ^ (Syntax.string_of_term_global thy (c $ t))) |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1889 |
| Sidecond t => Sidecond (c $ t)) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1890 |
| (c as Const (s, _), ts) => |
31877
e3de75d3b898
exported is_registered; added debug messages; removed extended parameter compilation in predicate compiler
bulwahn
parents:
31876
diff
changeset
|
1891 |
if is_registered thy s then |
31514 | 1892 |
let val (ts1, ts2) = chop (nparams_of thy s) ts |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1893 |
in Prem (ts2, list_comb (c, ts1)) end |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1894 |
else Sidecond t |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1895 |
| _ => Sidecond t) |
31514 | 1896 |
fun add_clause intr (clauses, arities) = |
1897 |
let |
|
1898 |
val _ $ t = Logic.strip_imp_concl intr; |
|
1899 |
val (Const (name, T), ts) = strip_comb t; |
|
1900 |
val (ts1, ts2) = chop nparams ts; |
|
1901 |
val prems = map (dest_prem o HOLogic.dest_Trueprop) (Logic.strip_imp_prems intr); |
|
1902 |
val (Ts, Us) = chop nparams (binder_types T) |
|
1903 |
in |
|
1904 |
(AList.update op = (name, these (AList.lookup op = clauses name) @ |
|
1905 |
[(ts2, prems)]) clauses, |
|
1906 |
AList.update op = (name, (map (fn U => (case strip_type U of |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1907 |
(Rs as _ :: _, Type ("bool", [])) => SOME (length Rs) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1908 |
| _ => NONE)) Ts, |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1909 |
length Us)) arities) |
31514 | 1910 |
end; |
1911 |
val (clauses, arities) = fold add_clause intrs ([], []); |
|
1912 |
in (preds, nparams, all_vs, param_vs, extra_modes, clauses, arities) end; |
|
1913 |
||
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1914 |
(** main function of predicate compiler **) |
31514 | 1915 |
|
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
1916 |
fun add_equations_of steps prednames thy = |
32317 | 1917 |
let |
1918 |
val _ = Output.tracing ("Starting predicate compiler for predicates " ^ commas prednames ^ "...") |
|
1919 |
val (preds, nparams, all_vs, param_vs, extra_modes, clauses, arities) = |
|
1920 |
prepare_intrs thy prednames |
|
1921 |
val _ = Output.tracing "Infering modes..." |
|
1922 |
val moded_clauses = #infer_modes steps thy extra_modes arities param_vs clauses |
|
1923 |
val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses |
|
1924 |
val _ = print_modes modes |
|
1925 |
val _ = print_moded_clauses thy moded_clauses |
|
1926 |
val _ = Output.tracing "Defining executable functions..." |
|
1927 |
val thy' = fold (#create_definitions steps preds) modes thy |
|
1928 |
|> Theory.checkpoint |
|
1929 |
val _ = Output.tracing "Compiling equations..." |
|
1930 |
val compiled_terms = |
|
1931 |
(#compile_preds steps) thy' all_vs param_vs preds moded_clauses |
|
1932 |
val _ = print_compiled_terms thy' compiled_terms |
|
1933 |
val _ = Output.tracing "Proving equations..." |
|
1934 |
val result_thms = #prove steps thy' clauses preds (extra_modes @ modes) |
|
1935 |
moded_clauses compiled_terms |
|
1936 |
val qname = #qname steps |
|
1937 |
(* val attrib = gn thy => Attrib.attribute_i thy Code.add_eqn_attrib *) |
|
1938 |
val attrib = fn thy => Attrib.attribute_i thy (Attrib.internal (K (Thm.declaration_attribute |
|
1939 |
(fn thm => Context.mapping (Code.add_eqn thm) I)))) |
|
1940 |
val thy'' = fold (fn (name, result_thms) => fn thy => snd (PureThy.add_thmss |
|
1941 |
[((Binding.qualify true (Long_Name.base_name name) (Binding.name qname), result_thms), |
|
1942 |
[attrib thy ])] thy)) |
|
1943 |
(maps_modes result_thms) thy' |
|
1944 |
|> Theory.checkpoint |
|
1945 |
in |
|
1946 |
thy'' |
|
1947 |
end |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1948 |
|
32314
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
1949 |
fun extend' value_of edges_of key (G, visited) = |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
1950 |
let |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
1951 |
val (G', v) = case try (Graph.get_node G) key of |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
1952 |
SOME v => (G, v) |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
1953 |
| NONE => (Graph.new_node (key, value_of key) G, value_of key) |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
1954 |
val (G'', visited') = fold (extend' value_of edges_of) (edges_of (key, v) \\ visited) |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
1955 |
(G', key :: visited) |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
1956 |
in |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
1957 |
(fold (Graph.add_edge o (pair key)) (edges_of (key, v)) G'', visited') |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
1958 |
end; |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
1959 |
|
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
1960 |
fun extend value_of edges_of key G = fst (extend' value_of edges_of key (G, [])) |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
1961 |
|
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1962 |
fun gen_add_equations steps names thy = |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1963 |
let |
32316
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
1964 |
val thy' = PredData.map (fold (extend (fetch_pred_data thy) (depending_preds_of thy)) names) thy |
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32315
diff
changeset
|
1965 |
|> Theory.checkpoint; |
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1966 |
fun strong_conn_of gr keys = |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1967 |
Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr) |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1968 |
val scc = strong_conn_of (PredData.get thy') names |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1969 |
val thy'' = fold_rev |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1970 |
(fn preds => fn thy => |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1971 |
if #are_not_defined steps thy preds then add_equations_of steps preds thy else thy) |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1972 |
scc thy' |> Theory.checkpoint |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1973 |
in thy'' end |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1974 |
|
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1975 |
(* different instantiantions of the predicate compiler *) |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1976 |
|
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1977 |
val add_equations = gen_add_equations |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1978 |
{infer_modes = infer_modes false, |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1979 |
create_definitions = create_definitions, |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1980 |
compile_preds = compile_preds PredicateCompFuns.compfuns mk_fun_of false, |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1981 |
prove = prove, |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1982 |
are_not_defined = (fn thy => forall (null o modes_of thy)), |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1983 |
qname = "equation"} |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1984 |
|
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1985 |
val add_sizelim_equations = gen_add_equations |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1986 |
{infer_modes = infer_modes false, |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1987 |
create_definitions = sizelim_create_definitions, |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1988 |
compile_preds = compile_preds PredicateCompFuns.compfuns mk_sizelim_fun_of true, |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1989 |
prove = prove_by_skip, |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1990 |
are_not_defined = (fn thy => fn preds => true), (* TODO *) |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1991 |
qname = "sizelim_equation" |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1992 |
} |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1993 |
|
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1994 |
val add_quickcheck_equations = gen_add_equations |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1995 |
{infer_modes = infer_modes_with_generator, |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1996 |
create_definitions = rpred_create_definitions, |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1997 |
compile_preds = compile_preds RPredCompFuns.compfuns mk_generator_of true, |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1998 |
prove = prove_by_skip, |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
1999 |
are_not_defined = (fn thy => fn preds => true), (* TODO *) |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
2000 |
qname = "rpred_equation"} |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
2001 |
|
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
2002 |
(** user interface **) |
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
2003 |
|
31106 | 2004 |
(* generation of case rules from user-given introduction rules *) |
2005 |
||
31550
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31549
diff
changeset
|
2006 |
fun mk_casesrule ctxt nparams introrules = |
31124 | 2007 |
let |
31573
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
2008 |
val intros = map (Logic.unvarify o prop_of) introrules |
31514 | 2009 |
val (pred, (params, args)) = strip_intro_concl nparams (hd intros) |
31106 | 2010 |
val ([propname], ctxt1) = Variable.variant_fixes ["thesis"] ctxt |
2011 |
val prop = HOLogic.mk_Trueprop (Free (propname, HOLogic.boolT)) |
|
2012 |
val (argnames, ctxt2) = Variable.variant_fixes |
|
2013 |
(map (fn i => "a" ^ string_of_int i) (1 upto (length args))) ctxt1 |
|
31514 | 2014 |
val argvs = map2 (curry Free) argnames (map fastype_of args) |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
2015 |
fun mk_case intro = |
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
2016 |
let |
31514 | 2017 |
val (_, (_, args)) = strip_intro_concl nparams intro |
31106 | 2018 |
val prems = Logic.strip_imp_prems intro |
2019 |
val eqprems = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (argvs ~~ args) |
|
2020 |
val frees = (fold o fold_aterms) |
|
2021 |
(fn t as Free _ => |
|
2022 |
if member (op aconv) params t then I else insert (op aconv) t |
|
2023 |
| _ => I) (args @ prems) [] |
|
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
2024 |
in fold Logic.all frees (Logic.list_implies (eqprems @ prems, prop)) end |
31106 | 2025 |
val assm = HOLogic.mk_Trueprop (list_comb (pred, params @ argvs)) |
2026 |
val cases = map mk_case intros |
|
31550
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31549
diff
changeset
|
2027 |
in Logic.list_implies (assm :: cases, prop) end; |
31106 | 2028 |
|
32312
26a9d0c69b8b
refactoring predicate compiler; added type synonyms and functions split_smode and split_mode
bulwahn
parents:
32311
diff
changeset
|
2029 |
(* code_pred_intro attribute *) |
32311
50f953140636
adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents:
32310
diff
changeset
|
2030 |
|
31876
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
2031 |
fun attrib f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I); |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
2032 |
|
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
2033 |
val code_pred_intros_attrib = attrib add_intro; |
9ab571673059
added diagnostic printing; changed proof for parameters; moved code
bulwahn
parents:
31784
diff
changeset
|
2034 |
|
31124 | 2035 |
local |
2036 |
||
31514 | 2037 |
(* TODO: make TheoryDataFun to GenericDataFun & remove duplication of local theory and theory *) |
2038 |
(* TODO: must create state to prove multiple cases *) |
|
31124 | 2039 |
fun generic_code_pred prep_const raw_const lthy = |
2040 |
let |
|
31550
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31549
diff
changeset
|
2041 |
val thy = ProofContext.theory_of lthy |
31124 | 2042 |
val const = prep_const thy raw_const |
32314
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
2043 |
val lthy' = LocalTheory.theory (PredData.map |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32313
diff
changeset
|
2044 |
(extend (fetch_pred_data thy) (depending_preds_of thy) const)) lthy |
31554
a6c6bf2751a0
added sporadic (Local)Theory.checkpoint, to enable parallel proof checking;
wenzelm
parents:
31550
diff
changeset
|
2045 |
|> LocalTheory.checkpoint |
31550
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31549
diff
changeset
|
2046 |
val thy' = ProofContext.theory_of lthy' |
31514 | 2047 |
val preds = Graph.all_preds (PredData.get thy') [const] |> filter_out (has_elim thy') |
31550
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31549
diff
changeset
|
2048 |
fun mk_cases const = |
31514 | 2049 |
let |
31550
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31549
diff
changeset
|
2050 |
val nparams = nparams_of thy' const |
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31549
diff
changeset
|
2051 |
val intros = intros_of thy' const |
31580 | 2052 |
in mk_casesrule lthy' nparams intros end |
31550
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31549
diff
changeset
|
2053 |
val cases_rules = map mk_cases preds |
31580 | 2054 |
val cases = |
2055 |
map (fn case_rule => RuleCases.Case {fixes = [], |
|
2056 |
assumes = [("", Logic.strip_imp_prems case_rule)], |
|
2057 |
binds = [], cases = []}) cases_rules |
|
2058 |
val case_env = map2 (fn p => fn c => (Long_Name.base_name p, SOME c)) preds cases |
|
32318
bca7fd849829
improved use of context with cases rule in predicate compiler; predicate compiler based on Main for faster debugging
bulwahn
parents:
32317
diff
changeset
|
2059 |
val lthy'' = lthy' |
bca7fd849829
improved use of context with cases rule in predicate compiler; predicate compiler based on Main for faster debugging
bulwahn
parents:
32317
diff
changeset
|
2060 |
|> fold Variable.auto_fixes cases_rules |
bca7fd849829
improved use of context with cases rule in predicate compiler; predicate compiler based on Main for faster debugging
bulwahn
parents:
32317
diff
changeset
|
2061 |
|> ProofContext.add_cases true case_env |
bca7fd849829
improved use of context with cases rule in predicate compiler; predicate compiler based on Main for faster debugging
bulwahn
parents:
32317
diff
changeset
|
2062 |
fun after_qed thms goal_ctxt = |
bca7fd849829
improved use of context with cases rule in predicate compiler; predicate compiler based on Main for faster debugging
bulwahn
parents:
32317
diff
changeset
|
2063 |
let |
bca7fd849829
improved use of context with cases rule in predicate compiler; predicate compiler based on Main for faster debugging
bulwahn
parents:
32317
diff
changeset
|
2064 |
val global_thms = ProofContext.export goal_ctxt |
bca7fd849829
improved use of context with cases rule in predicate compiler; predicate compiler based on Main for faster debugging
bulwahn
parents:
32317
diff
changeset
|
2065 |
(ProofContext.init (ProofContext.theory_of goal_ctxt)) (map the_single thms) |
bca7fd849829
improved use of context with cases rule in predicate compiler; predicate compiler based on Main for faster debugging
bulwahn
parents:
32317
diff
changeset
|
2066 |
in |
bca7fd849829
improved use of context with cases rule in predicate compiler; predicate compiler based on Main for faster debugging
bulwahn
parents:
32317
diff
changeset
|
2067 |
goal_ctxt |> LocalTheory.theory (fold set_elim global_thms #> add_equations [const]) |
bca7fd849829
improved use of context with cases rule in predicate compiler; predicate compiler based on Main for faster debugging
bulwahn
parents:
32317
diff
changeset
|
2068 |
end |
31124 | 2069 |
in |
31580 | 2070 |
Proof.theorem_i NONE after_qed (map (single o (rpair [])) cases_rules) lthy'' |
31124 | 2071 |
end; |
2072 |
||
2073 |
structure P = OuterParse |
|
31106 | 2074 |
|
31124 | 2075 |
in |
2076 |
||
2077 |
val code_pred = generic_code_pred (K I); |
|
31156 | 2078 |
val code_pred_cmd = generic_code_pred Code.read_const |
31124 | 2079 |
|
31514 | 2080 |
val setup = PredData.put (Graph.empty) #> |
2081 |
Attrib.setup @{binding code_pred_intros} (Scan.succeed (attrib add_intro)) |
|
2082 |
"adding alternative introduction rules for code generation of inductive predicates" |
|
2083 |
(* Attrib.setup @{binding code_ind_cases} (Scan.succeed add_elim_attrib) |
|
31124 | 2084 |
"adding alternative elimination rules for code generation of inductive predicates"; |
31514 | 2085 |
*) |
31124 | 2086 |
(*FIXME name discrepancy in attribs and ML code*) |
2087 |
(*FIXME intros should be better named intro*) |
|
31514 | 2088 |
(*FIXME why distinguished attribute for cases?*) |
31124 | 2089 |
|
2090 |
val _ = OuterSyntax.local_theory_to_proof "code_pred" |
|
2091 |
"prove equations for predicate specified by intro/elim rules" |
|
2092 |
OuterKeyword.thy_goal (P.term_group >> code_pred_cmd) |
|
2093 |
||
2094 |
end |
|
2095 |
||
2096 |
(*FIXME |
|
2097 |
- Naming of auxiliary rules necessary? |
|
31217 | 2098 |
- add default code equations P x y z = P_i_i_i x y z |
31124 | 2099 |
*) |
31106 | 2100 |
|
31169
f4c61cbea49d
added predicate transformation function for code generation
bulwahn
parents:
31107
diff
changeset
|
2101 |
(* transformation for code generation *) |
f4c61cbea49d
added predicate transformation function for code generation
bulwahn
parents:
31107
diff
changeset
|
2102 |
|
31217 | 2103 |
val eval_ref = ref (NONE : (unit -> term Predicate.pred) option); |
2104 |
||
32341
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2105 |
(*FIXME turn this into an LCF-guarded preprocessor for comprehensions*) |
31217 | 2106 |
fun analyze_compr thy t_compr = |
2107 |
let |
|
2108 |
val split = case t_compr of (Const (@{const_name Collect}, _) $ t) => t |
|
2109 |
| _ => error ("Not a set comprehension: " ^ Syntax.string_of_term_global thy t_compr); |
|
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32341
diff
changeset
|
2110 |
val (body, Ts, fp) = HOLogic.strip_psplits split; |
32341
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2111 |
val (pred as Const (name, T), all_args) = strip_comb body; |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2112 |
val (params, args) = chop (nparams_of thy name) all_args; |
31217 | 2113 |
val user_mode = map_filter I (map_index |
2114 |
(fn (i, t) => case t of Bound j => if j < length Ts then NONE |
|
32341
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2115 |
else SOME (i+1) | _ => SOME (i+1)) args); (*FIXME dangling bounds should not occur*) |
31217 | 2116 |
val modes = filter (fn Mode (_, is, _) => is = user_mode) |
31514 | 2117 |
(modes_of_term (all_modes_of thy) (list_comb (pred, params))); |
31217 | 2118 |
val m = case modes |
2119 |
of [] => error ("No mode possible for comprehension " |
|
2120 |
^ Syntax.string_of_term_global thy t_compr) |
|
2121 |
| [m] => m |
|
2122 |
| m :: _ :: _ => (warning ("Multiple modes possible for comprehension " |
|
2123 |
^ Syntax.string_of_term_global thy t_compr); m); |
|
32351 | 2124 |
val (inargs, outargs) = split_smode user_mode args; |
2125 |
val t_pred = list_comb (compile_expr NONE thy (m, list_comb (pred, params)), inargs); |
|
32341
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2126 |
val t_eval = if null outargs then t_pred else let |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2127 |
val outargs_bounds = map (fn Bound i => i) outargs; |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2128 |
val outargsTs = map (nth Ts) outargs_bounds; |
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32341
diff
changeset
|
2129 |
val T_pred = HOLogic.mk_tupleT outargsTs; |
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32341
diff
changeset
|
2130 |
val T_compr = HOLogic.mk_ptupleT fp Ts; |
32341
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2131 |
val arrange_bounds = map_index I outargs_bounds |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2132 |
|> sort (prod_ord (K EQUAL) int_ord) |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2133 |
|> map fst; |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2134 |
val arrange = funpow (length outargs_bounds - 1) HOLogic.mk_split |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2135 |
(Term.list_abs (map (pair "") outargsTs, |
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32341
diff
changeset
|
2136 |
HOLogic.mk_ptuple fp T_compr (map Bound arrange_bounds))) |
32351 | 2137 |
in mk_map PredicateCompFuns.compfuns T_pred T_compr arrange t_pred end |
31217 | 2138 |
in t_eval end; |
31169
f4c61cbea49d
added predicate transformation function for code generation
bulwahn
parents:
31107
diff
changeset
|
2139 |
|
32341
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2140 |
fun eval thy t_compr = |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2141 |
let |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2142 |
val t = analyze_compr thy t_compr; |
32351 | 2143 |
val T = dest_predT PredicateCompFuns.compfuns (fastype_of t); |
2144 |
val t' = mk_map PredicateCompFuns.compfuns T HOLogic.termT (HOLogic.term_of_const T) t; |
|
32341
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2145 |
in (T, Code_ML.eval NONE ("Predicate_Compile.eval_ref", eval_ref) Predicate.map thy t' []) end; |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2146 |
|
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2147 |
fun values ctxt k t_compr = |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2148 |
let |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2149 |
val thy = ProofContext.theory_of ctxt; |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2150 |
val (T, t) = eval thy t_compr; |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2151 |
val setT = HOLogic.mk_setT T; |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2152 |
val (ts, _) = Predicate.yieldn k t; |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2153 |
val elemsT = HOLogic.mk_set T ts; |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2154 |
in if k = ~1 orelse length ts < k then elemsT |
32683
7c1fe854ca6a
inter and union are mere abbreviations for inf and sup
haftmann
parents:
32351
diff
changeset
|
2155 |
else Const (@{const_name Lattices.sup}, setT --> setT --> setT) $ elemsT $ t_compr |
32341
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2156 |
end; |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2157 |
|
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2158 |
fun values_cmd modes k raw_t state = |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2159 |
let |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2160 |
val ctxt = Toplevel.context_of state; |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2161 |
val t = Syntax.read_term ctxt raw_t; |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2162 |
val t' = values ctxt k t; |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2163 |
val ty' = Term.type_of t'; |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2164 |
val ctxt' = Variable.auto_fixes t' ctxt; |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2165 |
val p = PrintMode.with_modes modes (fn () => |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2166 |
Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk, |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2167 |
Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) (); |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2168 |
in Pretty.writeln p end; |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2169 |
|
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2170 |
local structure P = OuterParse in |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2171 |
|
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2172 |
val opt_modes = Scan.optional (P.$$$ "(" |-- P.!!! (Scan.repeat1 P.xname --| P.$$$ ")")) []; |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2173 |
|
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2174 |
val _ = OuterSyntax.improper_command "values" "enumerate and print comprehensions" OuterKeyword.diag |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2175 |
(opt_modes -- Scan.optional P.nat ~1 -- P.term |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2176 |
>> (fn ((modes, k), t) => Toplevel.no_timing o Toplevel.keep |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2177 |
(values_cmd modes k t))); |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2178 |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
2179 |
end; |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
2180 |
|
32341
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2181 |
end; |
c8c17c2e6ceb
towards proper handling of argument order in comprehensions
haftmann
parents:
32287
diff
changeset
|
2182 |