author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Wed, 10 Feb 2010 12:30:26 +0100 | |
changeset 1121 | 8d3f92694e85 |
parent 1114 | 67dff6c1a123 |
child 1128 | 17ca92ab4660 |
permissions | -rw-r--r-- |
952
9c3b3eaecaff
use of equiv_relation_chk in quotient_term
Christian Urban <urbanc@in.tum.de>
parents:
884
diff
changeset
|
1 |
(* Title: quotient_def.thy |
9c3b3eaecaff
use of equiv_relation_chk in quotient_term
Christian Urban <urbanc@in.tum.de>
parents:
884
diff
changeset
|
2 |
Author: Cezary Kaliszyk and Christian Urban |
9c3b3eaecaff
use of equiv_relation_chk in quotient_term
Christian Urban <urbanc@in.tum.de>
parents:
884
diff
changeset
|
3 |
|
9c3b3eaecaff
use of equiv_relation_chk in quotient_term
Christian Urban <urbanc@in.tum.de>
parents:
884
diff
changeset
|
4 |
Definitions for constants on quotient types. |
9c3b3eaecaff
use of equiv_relation_chk in quotient_term
Christian Urban <urbanc@in.tum.de>
parents:
884
diff
changeset
|
5 |
|
9c3b3eaecaff
use of equiv_relation_chk in quotient_term
Christian Urban <urbanc@in.tum.de>
parents:
884
diff
changeset
|
6 |
*) |
277
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
7 |
|
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
8 |
signature QUOTIENT_DEF = |
774
b4ffb8826105
moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents:
768
diff
changeset
|
9 |
sig |
709 | 10 |
val quotient_def: mixfix -> Attrib.binding -> term -> term -> |
775 | 11 |
local_theory -> (term * thm) * local_theory |
789 | 12 |
|
705
f51c6069cd17
New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
670
diff
changeset
|
13 |
val quotdef_cmd: (Attrib.binding * string) * (mixfix * string) -> |
277
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
14 |
local_theory -> local_theory |
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
15 |
end; |
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
16 |
|
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
17 |
structure Quotient_Def: QUOTIENT_DEF = |
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
18 |
struct |
37636f2b1c19
separated the quotient_def into a separate file
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:
760
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:
768
diff
changeset
|
21 |
open Quotient_Term; |
762
baac4639ecef
avoided global "open"s - replaced by local "open"s
Christian Urban <urbanc@in.tum.de>
parents:
760
diff
changeset
|
22 |
|
279 | 23 |
(* wrapper for define *) |
277
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
24 |
fun define name mx attr rhs lthy = |
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
25 |
let |
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
26 |
val ((rhs, (_ , thm)), lthy') = |
331
345c422b1cb5
updated to Isabelle 22nd November
Christian Urban <urbanc@in.tum.de>
parents:
329
diff
changeset
|
27 |
Local_Theory.define ((name, mx), (attr, rhs)) lthy |
277
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
28 |
in |
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
29 |
((rhs, thm), lthy') |
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
30 |
end |
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
31 |
|
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
32 |
|
884
e49c6b6f37f4
tuned quotient_def.ML and cleaned somewhat LamEx.thy
Christian Urban <urbanc@in.tum.de>
parents:
869
diff
changeset
|
33 |
(** Interface and Syntax Setup **) |
850 | 34 |
|
884
e49c6b6f37f4
tuned quotient_def.ML and cleaned somewhat LamEx.thy
Christian Urban <urbanc@in.tum.de>
parents:
869
diff
changeset
|
35 |
(* The ML-interface for a quotient definition takes |
e49c6b6f37f4
tuned quotient_def.ML and cleaned somewhat LamEx.thy
Christian Urban <urbanc@in.tum.de>
parents:
869
diff
changeset
|
36 |
as argument: |
709 | 37 |
|
884
e49c6b6f37f4
tuned quotient_def.ML and cleaned somewhat LamEx.thy
Christian Urban <urbanc@in.tum.de>
parents:
869
diff
changeset
|
38 |
- the mixfix annotation |
e49c6b6f37f4
tuned quotient_def.ML and cleaned somewhat LamEx.thy
Christian Urban <urbanc@in.tum.de>
parents:
869
diff
changeset
|
39 |
- name and attributes |
e49c6b6f37f4
tuned quotient_def.ML and cleaned somewhat LamEx.thy
Christian Urban <urbanc@in.tum.de>
parents:
869
diff
changeset
|
40 |
- the new constant as term |
e49c6b6f37f4
tuned quotient_def.ML and cleaned somewhat LamEx.thy
Christian Urban <urbanc@in.tum.de>
parents:
869
diff
changeset
|
41 |
- the rhs of the definition as term |
850 | 42 |
|
884
e49c6b6f37f4
tuned quotient_def.ML and cleaned somewhat LamEx.thy
Christian Urban <urbanc@in.tum.de>
parents:
869
diff
changeset
|
43 |
It returns the defined constant and its definition |
e49c6b6f37f4
tuned quotient_def.ML and cleaned somewhat LamEx.thy
Christian Urban <urbanc@in.tum.de>
parents:
869
diff
changeset
|
44 |
theorem; stores the data in the qconsts data slot. |
e49c6b6f37f4
tuned quotient_def.ML and cleaned somewhat LamEx.thy
Christian Urban <urbanc@in.tum.de>
parents:
869
diff
changeset
|
45 |
|
e49c6b6f37f4
tuned quotient_def.ML and cleaned somewhat LamEx.thy
Christian Urban <urbanc@in.tum.de>
parents:
869
diff
changeset
|
46 |
Restriction: At the moment the right-hand side must |
e49c6b6f37f4
tuned quotient_def.ML and cleaned somewhat LamEx.thy
Christian Urban <urbanc@in.tum.de>
parents:
869
diff
changeset
|
47 |
be a terms composed of constant. Similarly the |
e49c6b6f37f4
tuned quotient_def.ML and cleaned somewhat LamEx.thy
Christian Urban <urbanc@in.tum.de>
parents:
869
diff
changeset
|
48 |
left-hand side must be a single constant. |
858 | 49 |
*) |
709 | 50 |
fun quotient_def mx attr lhs rhs lthy = |
293
653460d3e849
tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents:
290
diff
changeset
|
51 |
let |
884
e49c6b6f37f4
tuned quotient_def.ML and cleaned somewhat LamEx.thy
Christian Urban <urbanc@in.tum.de>
parents:
869
diff
changeset
|
52 |
val (lhs_str, lhs_ty) = dest_Free lhs handle TERM _ => error "Constant already defined." |
e49c6b6f37f4
tuned quotient_def.ML and cleaned somewhat LamEx.thy
Christian Urban <urbanc@in.tum.de>
parents:
869
diff
changeset
|
53 |
val _ = if null (strip_abs_vars rhs) then () else error "The definiens cannot be an abstraction" |
e49c6b6f37f4
tuned quotient_def.ML and cleaned somewhat LamEx.thy
Christian Urban <urbanc@in.tum.de>
parents:
869
diff
changeset
|
54 |
|
e49c6b6f37f4
tuned quotient_def.ML and cleaned somewhat LamEx.thy
Christian Urban <urbanc@in.tum.de>
parents:
869
diff
changeset
|
55 |
val qconst_bname = Binding.name lhs_str |
1097
551eacf071d7
More indentation, names and todo cleaning in the quotient package
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
952
diff
changeset
|
56 |
val absrep_trm = absrep_fun AbsF lthy (fastype_of rhs, lhs_ty) $ rhs |
854
5961edda27d7
absrep_fun and equiv_relation do not produce anymore spurious maps; two problems arose in IntEx, which are marked with "INJECTION PROBLEM"
Christian Urban <urbanc@in.tum.de>
parents:
833
diff
changeset
|
57 |
val prop = Logic.mk_equals (lhs, Syntax.check_term lthy absrep_trm) |
663
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
598
diff
changeset
|
58 |
val (_, prop') = LocalDefs.cert_def lthy prop |
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
598
diff
changeset
|
59 |
val (_, newrhs) = Primitive_Defs.abs_def prop' |
310
fec6301a1989
added a container for quotient constants (does not work yet though)
Christian Urban <urbanc@in.tum.de>
parents:
307
diff
changeset
|
60 |
|
663
0dd10a900cae
Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
598
diff
changeset
|
61 |
val ((trm, thm), lthy') = define qconst_bname mx attr newrhs lthy |
310
fec6301a1989
added a container for quotient constants (does not work yet though)
Christian Urban <urbanc@in.tum.de>
parents:
307
diff
changeset
|
62 |
|
800 | 63 |
(* data storage *) |
799
0755f8fd56b3
renamed transfer to transform (Markus)
Christian Urban <urbanc@in.tum.de>
parents:
789
diff
changeset
|
64 |
fun qcinfo phi = transform_qconsts phi {qconst = trm, rconst = rhs, def = thm} |
869
ce5f78f0eac5
Finished organising an efficient datastructure for qconst_info.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
868
diff
changeset
|
65 |
fun trans_name phi = (fst o dest_Const o #qconst) (qcinfo phi) |
319 | 66 |
val lthy'' = Local_Theory.declaration true |
869
ce5f78f0eac5
Finished organising an efficient datastructure for qconst_info.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
868
diff
changeset
|
67 |
(fn phi => qconsts_update_gen (trans_name phi) (qcinfo phi)) lthy' |
293
653460d3e849
tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents:
290
diff
changeset
|
68 |
in |
310
fec6301a1989
added a container for quotient constants (does not work yet though)
Christian Urban <urbanc@in.tum.de>
parents:
307
diff
changeset
|
69 |
((trm, thm), lthy'') |
293
653460d3e849
tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents:
290
diff
changeset
|
70 |
end |
653460d3e849
tuned the code in quotient and quotient_def
Christian Urban <urbanc@in.tum.de>
parents:
290
diff
changeset
|
71 |
|
804 | 72 |
fun quotdef_cmd ((attr, lhs_str), (mx, rhs_str)) lthy = |
277
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
73 |
let |
804 | 74 |
val lhs = Syntax.read_term lthy lhs_str |
75 |
val rhs = Syntax.read_term lthy rhs_str |
|
1114 | 76 |
val lthy' = Variable.declare_term lhs lthy |
77 |
val lthy'' = Variable.declare_term rhs lthy' |
|
277
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
78 |
in |
1114 | 79 |
quotient_def mx attr lhs rhs lthy'' |> snd |
277
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
80 |
end |
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
81 |
|
705
f51c6069cd17
New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
670
diff
changeset
|
82 |
val _ = OuterKeyword.keyword "as"; |
f51c6069cd17
New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
670
diff
changeset
|
83 |
|
279 | 84 |
val quotdef_parser = |
705
f51c6069cd17
New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
670
diff
changeset
|
85 |
(SpecParse.opt_thm_name ":" -- |
760
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
719
diff
changeset
|
86 |
OuterParse.term) -- |
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
719
diff
changeset
|
87 |
(OuterParse.opt_mixfix' --| OuterParse.$$$ "as" -- |
c1989de100b4
various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents:
719
diff
changeset
|
88 |
OuterParse.term) |
279 | 89 |
|
767
37285ec4387d
on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
90 |
val _ = OuterSyntax.local_theory "quotient_definition" |
37285ec4387d
on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
91 |
"definition for constants over the quotient type" |
37285ec4387d
on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents:
762
diff
changeset
|
92 |
OuterKeyword.thy_decl (quotdef_parser >> quotdef_cmd) |
277
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
93 |
|
37636f2b1c19
separated the quotient_def into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
94 |
end; (* structure *) |