author | Christian Urban <urbanc@in.tum.de> |
Wed, 23 Dec 2009 21:30:23 +0100 | |
changeset 782 | 86c7ed9f354f |
parent 779 | 3b21b24a5fb6 |
child 783 | 06e17083e90b |
permissions | -rw-r--r-- |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1 |
signature QUOTIENT_TERM = |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
2 |
sig |
774
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
3 |
exception LIFT_MATCH of string |
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
4 |
|
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
5 |
datatype flag = absF | repF |
776
d1064fa29424
renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Christian Urban <urbanc@in.tum.de>
parents:
775
diff
changeset
|
6 |
|
d1064fa29424
renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Christian Urban <urbanc@in.tum.de>
parents:
775
diff
changeset
|
7 |
val absrep_fun: flag -> Proof.context -> (typ * typ) -> term |
d1064fa29424
renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Christian Urban <urbanc@in.tum.de>
parents:
775
diff
changeset
|
8 |
val absrep_fun_chk: flag -> Proof.context -> (typ * typ) -> term |
774
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
9 |
|
775 | 10 |
val regularize_trm: Proof.context -> (term * term) -> term |
776
d1064fa29424
renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Christian Urban <urbanc@in.tum.de>
parents:
775
diff
changeset
|
11 |
val regularize_trm_chk: Proof.context -> (term * term) -> term |
d1064fa29424
renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Christian Urban <urbanc@in.tum.de>
parents:
775
diff
changeset
|
12 |
|
774
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
13 |
val inj_repabs_trm: Proof.context -> (term * term) -> term |
776
d1064fa29424
renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Christian Urban <urbanc@in.tum.de>
parents:
775
diff
changeset
|
14 |
val inj_repabs_trm_chk: Proof.context -> (term * term) -> term |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
15 |
end; |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
16 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
17 |
structure Quotient_Term: QUOTIENT_TERM = |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
18 |
struct |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
19 |
|
762
baac4639ecef
avoided global "open"s - replaced by local "open"s
Christian Urban <urbanc@in.tum.de>
parents:
761
diff
changeset
|
20 |
open Quotient_Info; |
774
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
21 |
|
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
22 |
exception LIFT_MATCH of string |
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
23 |
|
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
24 |
(*******************************) |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
25 |
(* Aggregate Rep/Abs Functions *) |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
26 |
(*******************************) |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
27 |
|
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
28 |
(* The flag repF is for types in negative position, while absF is *) |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
29 |
(* for types in positive position. Because of this, function types *) |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
30 |
(* need to be treated specially, since there the polarity changes. *) |
774
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
31 |
|
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
32 |
datatype flag = absF | repF |
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
33 |
|
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
34 |
fun negF absF = repF |
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
35 |
| negF repF = absF |
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
36 |
|
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
37 |
fun mk_identity ty = Const (@{const_name "id"}, ty --> ty) |
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
38 |
|
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
39 |
fun mk_compose flag (trm1, trm2) = |
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
40 |
case flag of |
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
41 |
absF => Const (@{const_name "comp"}, dummyT) $ trm1 $ trm2 |
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
42 |
| repF => Const (@{const_name "comp"}, dummyT) $ trm2 $ trm1 |
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
43 |
|
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
44 |
fun get_map ctxt ty_str = |
774
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
45 |
let |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
46 |
val thy = ProofContext.theory_of ctxt |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
47 |
val exc = LIFT_MATCH (space_implode " " ["absrep_fun: no map for type", quote ty_str, "."]) |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
48 |
val info = maps_lookup thy ty_str handle NotFound => raise exc |
774
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
49 |
in |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
50 |
Const (#mapfun info, dummyT) |
774
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
51 |
end |
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
52 |
|
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
53 |
fun get_absrep_const flag ctxt _ qty = |
774
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
54 |
(* FIXME: check here that the type-constructors of _ and qty are related *) |
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
55 |
let |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
56 |
val thy = ProofContext.theory_of ctxt |
774
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
57 |
val qty_name = Long_Name.base_name (fst (dest_Type qty)) |
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
58 |
in |
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
59 |
case flag of |
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
60 |
absF => Const (Sign.full_bname thy ("abs_" ^ qty_name), dummyT) |
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
61 |
| repF => Const (Sign.full_bname thy ("rep_" ^ qty_name), dummyT) |
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
62 |
end |
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
63 |
|
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
64 |
fun absrep_fun flag ctxt (rty, qty) = |
774
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
65 |
if rty = qty then mk_identity qty else |
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
66 |
case (rty, qty) of |
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
67 |
(Type ("fun", [ty1, ty2]), Type ("fun", [ty1', ty2'])) => |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
68 |
let |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
69 |
val arg1 = absrep_fun (negF flag) ctxt (ty1, ty1') |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
70 |
val arg2 = absrep_fun flag ctxt (ty2, ty2') |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
71 |
in |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
72 |
list_comb (get_map ctxt "fun", [arg1, arg2]) |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
73 |
end |
774
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
74 |
| (Type (s, _), Type (s', [])) => |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
75 |
if s = s' |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
76 |
then mk_identity qty |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
77 |
else get_absrep_const flag ctxt rty qty |
774
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
78 |
| (Type (s, tys), Type (s', tys')) => |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
79 |
let |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
80 |
val args = map (absrep_fun flag ctxt) (tys ~~ tys') |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
81 |
val result = list_comb (get_map ctxt s, args) |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
82 |
in |
774
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
83 |
if s = s' |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
84 |
then result |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
85 |
else mk_compose flag (get_absrep_const flag ctxt rty qty, result) |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
86 |
end |
774
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
87 |
| (TFree x, TFree x') => |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
88 |
if x = x' |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
89 |
then mk_identity qty |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
90 |
else raise (LIFT_MATCH "absrep_fun (frees)") |
776
d1064fa29424
renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Christian Urban <urbanc@in.tum.de>
parents:
775
diff
changeset
|
91 |
| (TVar _, TVar _) => raise (LIFT_MATCH "absrep_fun (vars)") |
d1064fa29424
renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Christian Urban <urbanc@in.tum.de>
parents:
775
diff
changeset
|
92 |
| _ => raise (LIFT_MATCH "absrep_fun (default)") |
774
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
93 |
|
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
94 |
fun absrep_fun_chk flag ctxt (rty, qty) = |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
95 |
absrep_fun flag ctxt (rty, qty) |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
96 |
|> Syntax.check_term ctxt |
762
baac4639ecef
avoided global "open"s - replaced by local "open"s
Christian Urban <urbanc@in.tum.de>
parents:
761
diff
changeset
|
97 |
|
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
98 |
|
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
99 |
(* Regularizing an rtrm means: |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
100 |
|
760
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
101 |
- Quantifiers over types that need lifting are replaced |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
102 |
by bounded quantifiers, for example: |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
103 |
|
760
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
104 |
All P ----> All (Respects R) P |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
105 |
|
760
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
106 |
where the aggregate relation R is given by the rty and qty; |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
107 |
|
760
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
108 |
- Abstractions over types that need lifting are replaced |
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
109 |
by bounded abstractions, for example: |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
110 |
|
760
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
111 |
%x. P ----> Ball (Respects R) %x. P |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
112 |
|
760
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
113 |
- Equalities over types that need lifting are replaced by |
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
114 |
corresponding equivalence relations, for example: |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
115 |
|
760
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
116 |
A = B ----> R A B |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
117 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
118 |
or |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
119 |
|
760
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
120 |
A = B ----> (R ===> R) A B |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
121 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
122 |
for more complicated types of A and B |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
123 |
*) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
124 |
|
760
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
125 |
(* builds the aggregate equivalence relation *) |
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
126 |
(* that will be the argument of Respects *) |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
127 |
fun mk_resp_arg ctxt (rty, qty) = |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
128 |
let |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
129 |
val thy = ProofContext.theory_of ctxt |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
130 |
in |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
131 |
if rty = qty |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
132 |
then HOLogic.eq_const rty |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
133 |
else |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
134 |
case (rty, qty) of |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
135 |
(Type (s, tys), Type (s', tys')) => |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
136 |
if s = s' |
761 | 137 |
then |
138 |
let |
|
139 |
val exc = LIFT_MATCH ("mk_resp_arg (no map function found for type " ^ s) |
|
140 |
val map_info = maps_lookup thy s handle NotFound => raise exc |
|
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
141 |
val args = map (mk_resp_arg ctxt) (tys ~~ tys') |
761 | 142 |
in |
143 |
list_comb (Const (#relfun map_info, dummyT), args) |
|
144 |
end |
|
145 |
else |
|
146 |
let |
|
147 |
val SOME qinfo = quotdata_lookup_thy thy s' |
|
148 |
(* FIXME: check in this case that the rty and qty *) |
|
149 |
(* FIXME: correspond to each other *) |
|
150 |
val (s, _) = dest_Const (#rel qinfo) |
|
151 |
(* FIXME: the relation should only be the string *) |
|
152 |
(* FIXME: and the type needs to be calculated as below; *) |
|
153 |
(* FIXME: maybe one should actually have a term *) |
|
154 |
(* FIXME: and one needs to force it to have this type *) |
|
155 |
in |
|
156 |
Const (s, rty --> rty --> @{typ bool}) |
|
157 |
end |
|
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
158 |
| _ => HOLogic.eq_const dummyT |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
159 |
(* FIXME: check that the types correspond to each other? *) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
160 |
end |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
161 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
162 |
val mk_babs = Const (@{const_name Babs}, dummyT) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
163 |
val mk_ball = Const (@{const_name Ball}, dummyT) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
164 |
val mk_bex = Const (@{const_name Bex}, dummyT) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
165 |
val mk_resp = Const (@{const_name Respects}, dummyT) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
166 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
167 |
(* - applies f to the subterm of an abstraction, *) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
168 |
(* otherwise to the given term, *) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
169 |
(* - used by regularize, therefore abstracted *) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
170 |
(* variables do not have to be treated specially *) |
775 | 171 |
fun apply_subt f (trm1, trm2) = |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
172 |
case (trm1, trm2) of |
775 | 173 |
(Abs (x, T, t), Abs (_ , _, t')) => Abs (x, T, f (t, t')) |
174 |
| _ => f (trm1, trm2) |
|
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
175 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
176 |
(* the major type of All and Ex quantifiers *) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
177 |
fun qnt_typ ty = domain_type (domain_type ty) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
178 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
179 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
180 |
(* produces a regularized version of rtrm *) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
181 |
(* *) |
760
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
182 |
(* - the result still contains dummyTs *) |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
183 |
(* *) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
184 |
(* - for regularisation we do not need any *) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
185 |
(* special treatment of bound variables *) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
186 |
|
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
187 |
fun regularize_trm ctxt (rtrm, qtrm) = |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
188 |
case (rtrm, qtrm) of |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
189 |
(Abs (x, ty, t), Abs (_, ty', t')) => |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
190 |
let |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
191 |
val subtrm = Abs(x, ty, regularize_trm ctxt (t, t')) |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
192 |
in |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
193 |
if ty = ty' then subtrm |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
194 |
else mk_babs $ (mk_resp $ mk_resp_arg ctxt (ty, ty')) $ subtrm |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
195 |
end |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
196 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
197 |
| (Const (@{const_name "All"}, ty) $ t, Const (@{const_name "All"}, ty') $ t') => |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
198 |
let |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
199 |
val subtrm = apply_subt (regularize_trm ctxt) (t, t') |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
200 |
in |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
201 |
if ty = ty' then Const (@{const_name "All"}, ty) $ subtrm |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
202 |
else mk_ball $ (mk_resp $ mk_resp_arg ctxt (qnt_typ ty, qnt_typ ty')) $ subtrm |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
203 |
end |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
204 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
205 |
| (Const (@{const_name "Ex"}, ty) $ t, Const (@{const_name "Ex"}, ty') $ t') => |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
206 |
let |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
207 |
val subtrm = apply_subt (regularize_trm ctxt) (t, t') |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
208 |
in |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
209 |
if ty = ty' then Const (@{const_name "Ex"}, ty) $ subtrm |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
210 |
else mk_bex $ (mk_resp $ mk_resp_arg ctxt (qnt_typ ty, qnt_typ ty')) $ subtrm |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
211 |
end |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
212 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
213 |
| (* equalities need to be replaced by appropriate equivalence relations *) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
214 |
(Const (@{const_name "op ="}, ty), Const (@{const_name "op ="}, ty')) => |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
215 |
if ty = ty' then rtrm |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
216 |
else mk_resp_arg ctxt (domain_type ty, domain_type ty') |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
217 |
|
760
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
218 |
| (* in this case we just check whether the given equivalence relation is correct *) |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
219 |
(rel, Const (@{const_name "op ="}, ty')) => |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
220 |
let |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
221 |
val exc = LIFT_MATCH "regularise (relation mismatch)" |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
222 |
val rel_ty = (fastype_of rel) handle TERM _ => raise exc |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
223 |
val rel' = mk_resp_arg ctxt (domain_type rel_ty, domain_type ty') |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
224 |
in |
760
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
225 |
if rel' aconv rel then rtrm else raise exc |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
226 |
end |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
227 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
228 |
| (_, Const _) => |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
229 |
let |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
230 |
fun same_name (Const (s, T)) (Const (s', T')) = (s = s') (*andalso (T = T')*) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
231 |
| same_name _ _ = false |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
232 |
(* TODO/FIXME: This test is not enough. *) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
233 |
(* Why? *) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
234 |
(* Because constants can have the same name but not be the same |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
235 |
constant. All overloaded constants have the same name but because |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
236 |
of different types they do differ. |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
237 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
238 |
This code will let one write a theorem where plus on nat is |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
239 |
matched to plus on int, even if the latter is defined differently. |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
240 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
241 |
This would result in hard to understand failures in injection and |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
242 |
cleaning. *) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
243 |
(* cu: if I also test the type, then something else breaks *) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
244 |
in |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
245 |
if same_name rtrm qtrm then rtrm |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
246 |
else |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
247 |
let |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
248 |
val thy = ProofContext.theory_of ctxt |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
249 |
val qtrm_str = Syntax.string_of_term ctxt qtrm |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
250 |
val exc1 = LIFT_MATCH ("regularize (constant " ^ qtrm_str ^ " not found)") |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
251 |
val exc2 = LIFT_MATCH ("regularize (constant " ^ qtrm_str ^ " mismatch)") |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
252 |
val rtrm' = (#rconst (qconsts_lookup thy qtrm)) handle NotFound => raise exc1 |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
253 |
in |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
254 |
if Pattern.matches thy (rtrm', rtrm) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
255 |
then rtrm else raise exc2 |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
256 |
end |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
257 |
end |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
258 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
259 |
| (t1 $ t2, t1' $ t2') => |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
260 |
(regularize_trm ctxt (t1, t1')) $ (regularize_trm ctxt (t2, t2')) |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
261 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
262 |
| (Bound i, Bound i') => |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
263 |
if i = i' then rtrm |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
264 |
else raise (LIFT_MATCH "regularize (bounds mismatch)") |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
265 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
266 |
| _ => |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
267 |
let |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
268 |
val rtrm_str = Syntax.string_of_term ctxt rtrm |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
269 |
val qtrm_str = Syntax.string_of_term ctxt qtrm |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
270 |
in |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
271 |
raise (LIFT_MATCH ("regularize failed (default: " ^ rtrm_str ^ "," ^ qtrm_str ^ ")")) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
272 |
end |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
273 |
|
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
274 |
fun regularize_trm_chk ctxt (rtrm, qtrm) = |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
275 |
regularize_trm ctxt (rtrm, qtrm) |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
276 |
|> Syntax.check_term ctxt |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
277 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
278 |
(* |
760
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
279 |
Injection of Rep/Abs means: |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
280 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
281 |
For abstractions |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
282 |
: |
760
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
283 |
* If the type of the abstraction needs lifting, then we add Rep/Abs |
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
284 |
around the abstraction; otherwise we leave it unchanged. |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
285 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
286 |
For applications: |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
287 |
|
760
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
288 |
* If the application involves a bounded quantifier, we recurse on |
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
289 |
the second argument. If the application is a bounded abstraction, |
782
86c7ed9f354f
cleaed a bit function mk_typedef_main
Christian Urban <urbanc@in.tum.de>
parents:
779
diff
changeset
|
290 |
we always put an Rep/Abs around it (since bounded abstractions |
86c7ed9f354f
cleaed a bit function mk_typedef_main
Christian Urban <urbanc@in.tum.de>
parents:
779
diff
changeset
|
291 |
are assumed to always need lifting). Otherwise we recurse on both |
86c7ed9f354f
cleaed a bit function mk_typedef_main
Christian Urban <urbanc@in.tum.de>
parents:
779
diff
changeset
|
292 |
arguments. |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
293 |
|
760
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
294 |
For constants: |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
295 |
|
760
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
296 |
* If the constant is (op =), we leave it always unchanged. |
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
297 |
Otherwise the type of the constant needs lifting, we put |
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
298 |
and Rep/Abs around it. |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
299 |
|
760
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
300 |
For free variables: |
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
301 |
|
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
302 |
* We put aRep/Abs around it if the type needs lifting. |
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
303 |
|
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
304 |
Vars case cannot occur. |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
305 |
*) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
306 |
|
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
307 |
fun mk_repabs ctxt (T, T') trm = |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
308 |
absrep_fun repF ctxt (T, T') $ (absrep_fun absF ctxt (T, T') $ trm) |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
309 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
310 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
311 |
(* bound variables need to be treated properly, *) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
312 |
(* as the type of subterms needs to be calculated *) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
313 |
|
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
314 |
fun inj_repabs_trm ctxt (rtrm, qtrm) = |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
315 |
case (rtrm, qtrm) of |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
316 |
(Const (@{const_name "Ball"}, T) $ r $ t, Const (@{const_name "All"}, _) $ t') => |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
317 |
Const (@{const_name "Ball"}, T) $ r $ (inj_repabs_trm ctxt (t, t')) |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
318 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
319 |
| (Const (@{const_name "Bex"}, T) $ r $ t, Const (@{const_name "Ex"}, _) $ t') => |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
320 |
Const (@{const_name "Bex"}, T) $ r $ (inj_repabs_trm ctxt (t, t')) |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
321 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
322 |
| (Const (@{const_name "Babs"}, T) $ r $ t, t' as (Abs _)) => |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
323 |
let |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
324 |
val rty = fastype_of rtrm |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
325 |
val qty = fastype_of qtrm |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
326 |
in |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
327 |
mk_repabs ctxt (rty, qty) (Const (@{const_name "Babs"}, T) $ r $ (inj_repabs_trm ctxt (t, t'))) |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
328 |
end |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
329 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
330 |
| (Abs (x, T, t), Abs (x', T', t')) => |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
331 |
let |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
332 |
val rty = fastype_of rtrm |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
333 |
val qty = fastype_of qtrm |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
334 |
val (y, s) = Term.dest_abs (x, T, t) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
335 |
val (_, s') = Term.dest_abs (x', T', t') |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
336 |
val yvar = Free (y, T) |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
337 |
val result = Term.lambda_name (y, yvar) (inj_repabs_trm ctxt (s, s')) |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
338 |
in |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
339 |
if rty = qty then result |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
340 |
else mk_repabs ctxt (rty, qty) result |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
341 |
end |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
342 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
343 |
| (t $ s, t' $ s') => |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
344 |
(inj_repabs_trm ctxt (t, t')) $ (inj_repabs_trm ctxt (s, s')) |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
345 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
346 |
| (Free (_, T), Free (_, T')) => |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
347 |
if T = T' then rtrm |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
348 |
else mk_repabs ctxt (T, T') rtrm |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
349 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
350 |
| (_, Const (@{const_name "op ="}, _)) => rtrm |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
351 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
352 |
| (_, Const (_, T')) => |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
353 |
let |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
354 |
val rty = fastype_of rtrm |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
355 |
in |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
356 |
if rty = T' then rtrm |
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
357 |
else mk_repabs ctxt (rty, T') rtrm |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
358 |
end |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
359 |
|
760
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
758
diff
changeset
|
360 |
| _ => raise (LIFT_MATCH "injection (default)") |
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
361 |
|
779
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
362 |
fun inj_repabs_trm_chk ctxt (rtrm, qtrm) = |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
363 |
inj_repabs_trm ctxt (rtrm, qtrm) |
3b21b24a5fb6
corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
776
diff
changeset
|
364 |
|> Syntax.check_term ctxt |
776
d1064fa29424
renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Christian Urban <urbanc@in.tum.de>
parents:
775
diff
changeset
|
365 |
|
758
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
366 |
end; (* structure *) |
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
367 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
368 |
|
3104d62e7a16
moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
369 |