author | wenzelm |
Mon, 07 Nov 2005 19:03:02 +0100 | |
changeset 18109 | 94b528311e22 |
parent 17833 | 8631dfe017a8 |
child 18358 | 0a733e11021a |
permissions | -rw-r--r-- |
16696 | 1 |
(* Title: HOLCF/pcpodef_package.ML |
2 |
ID: $Id$ |
|
3 |
Author: Brian Huffman |
|
4 |
||
5 |
Gordon/HOL-style type definitions for HOLCF. |
|
6 |
*) |
|
7 |
||
8 |
signature PCPODEF_PACKAGE = |
|
9 |
sig |
|
10 |
val quiet_mode: bool ref |
|
11 |
val pcpodef_proof: (bool * string) * (bstring * string list * mixfix) * string |
|
17336 | 12 |
* (string * string) option -> theory -> Proof.state |
16696 | 13 |
val pcpodef_proof_i: (bool * string) * (bstring * string list * mixfix) * term |
17336 | 14 |
* (string * string) option -> theory -> Proof.state |
16696 | 15 |
val cpodef_proof: (bool * string) * (bstring * string list * mixfix) * string |
17336 | 16 |
* (string * string) option -> theory -> Proof.state |
16696 | 17 |
val cpodef_proof_i: (bool * string) * (bstring * string list * mixfix) * term |
17336 | 18 |
* (string * string) option -> theory -> Proof.state |
16696 | 19 |
end; |
20 |
||
21 |
structure PcpodefPackage: PCPODEF_PACKAGE = |
|
22 |
struct |
|
23 |
||
24 |
||
25 |
(** theory context references **) |
|
26 |
||
27 |
val typedef_po = thm "typedef_po"; |
|
28 |
val typedef_cpo = thm "typedef_cpo"; |
|
16919
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
29 |
val typedef_pcpo = thm "typedef_pcpo"; |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
30 |
val typedef_lub = thm "typedef_lub"; |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
31 |
val typedef_thelub = thm "typedef_thelub"; |
17833 | 32 |
val typedef_compact = thm "typedef_compact"; |
16696 | 33 |
val cont_Rep = thm "typedef_cont_Rep"; |
34 |
val cont_Abs = thm "typedef_cont_Abs"; |
|
35 |
val Rep_strict = thm "typedef_Rep_strict"; |
|
36 |
val Abs_strict = thm "typedef_Abs_strict"; |
|
37 |
val Rep_defined = thm "typedef_Rep_defined"; |
|
38 |
val Abs_defined = thm "typedef_Abs_defined"; |
|
39 |
||
40 |
||
41 |
(** type definitions **) |
|
42 |
||
43 |
(* messages *) |
|
44 |
||
45 |
val quiet_mode = ref false; |
|
46 |
fun message s = if ! quiet_mode then () else writeln s; |
|
47 |
||
48 |
||
49 |
(* prepare_cpodef *) |
|
50 |
||
51 |
fun read_term thy used s = |
|
52 |
#1 (Thm.read_def_cterm (thy, K NONE, K NONE) used true (s, HOLogic.typeT)); |
|
53 |
||
54 |
fun cert_term thy _ t = Thm.cterm_of thy t handle TERM (msg, _) => error msg; |
|
55 |
||
56 |
fun err_in_cpodef name = |
|
57 |
error ("The error(s) above occurred in cpodef " ^ quote name); |
|
58 |
||
59 |
fun adm_const T = Const ("Adm.adm", (T --> HOLogic.boolT) --> HOLogic.boolT); |
|
60 |
fun mk_adm (x, T, P) = adm_const T $ absfree (x, T, P); |
|
61 |
||
62 |
fun gen_prepare_pcpodef prep_term pcpo def name (t, vs, mx) raw_set opt_morphs thy = |
|
63 |
let |
|
64 |
val full = Sign.full_name thy; |
|
65 |
||
66 |
(*rhs*) |
|
67 |
val full_name = full name; |
|
68 |
val cset = prep_term thy vs raw_set; |
|
69 |
val {T = setT, t = set, ...} = Thm.rep_cterm cset; |
|
70 |
val rhs_tfrees = term_tfrees set; |
|
71 |
val oldT = HOLogic.dest_setT setT handle TYPE _ => |
|
72 |
error ("Not a set type: " ^ quote (Sign.string_of_typ thy setT)); |
|
73 |
fun mk_nonempty A = |
|
74 |
HOLogic.mk_exists ("x", oldT, HOLogic.mk_mem (Free ("x", oldT), A)); |
|
75 |
fun mk_admissible A = |
|
76 |
mk_adm ("x", oldT, HOLogic.mk_mem (Free ("x", oldT), A)); |
|
77 |
fun mk_UU_mem A = HOLogic.mk_mem (Const ("Pcpo.UU", oldT), A); |
|
78 |
val goal = if pcpo |
|
79 |
then HOLogic.mk_Trueprop (HOLogic.mk_conj (mk_UU_mem set, mk_admissible set)) |
|
80 |
else HOLogic.mk_Trueprop (HOLogic.mk_conj (mk_nonempty set, mk_admissible set)); |
|
81 |
||
82 |
(*lhs*) |
|
17325 | 83 |
val lhs_tfrees = map (fn v => (v, AList.lookup (op =) rhs_tfrees v |> the_default HOLogic.typeS)) vs; |
16696 | 84 |
val lhs_sorts = map snd lhs_tfrees; |
85 |
val tname = Syntax.type_name t mx; |
|
86 |
val full_tname = full tname; |
|
87 |
val newT = Type (full_tname, map TFree lhs_tfrees); |
|
88 |
||
89 |
val (Rep_name, Abs_name) = getOpt (opt_morphs, ("Rep_" ^ name, "Abs_" ^ name)); |
|
90 |
val RepC = Const (full Rep_name, newT --> oldT); |
|
17810
3bdf516d93d8
cleaned up; renamed "Porder.op <<" to "Porder.<<"
huffman
parents:
17336
diff
changeset
|
91 |
fun lessC T = Const ("Porder.<<", T --> T --> HOLogic.boolT); |
16696 | 92 |
val less_def = ("less_" ^ name ^ "_def", Logic.mk_equals (lessC newT, |
93 |
Abs ("x", newT, Abs ("y", newT, lessC oldT $ (RepC $ Bound 1) $ (RepC $ Bound 0))))); |
|
94 |
||
95 |
fun option_fold_rule NONE = I |
|
96 |
| option_fold_rule (SOME def) = fold_rule [def]; |
|
97 |
||
98 |
fun make_po tac theory = |
|
99 |
theory |
|
100 |
|> TypedefPackage.add_typedef_i def (SOME name) (t, vs, mx) set opt_morphs tac |
|
101 |
|>> AxClass.add_inst_arity_i (full_tname, lhs_sorts, ["Porder.sq_ord"]) |
|
102 |
(AxClass.intro_classes_tac []) |
|
103 |
|>>> PureThy.add_defs_i true [Thm.no_attributes less_def] |
|
104 |
|> (fn (theory', ({type_definition, set_def, ...}, [less_definition])) => |
|
105 |
theory' |
|
106 |
|> AxClass.add_inst_arity_i (full_tname, lhs_sorts, ["Porder.po"]) |
|
107 |
(Tactic.rtac (typedef_po OF [type_definition, less_definition]) 1) |
|
108 |
|> rpair (type_definition, less_definition, set_def)); |
|
109 |
||
16919
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
110 |
fun make_cpo admissible (theory, defs as (type_def, less_def, set_def)) = |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
111 |
let |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
112 |
val admissible' = option_fold_rule set_def admissible; |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
113 |
val cpo_thms = [type_def, less_def, admissible']; |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
114 |
val (theory', _) = |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
115 |
theory |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
116 |
|> AxClass.add_inst_arity_i (full_tname, lhs_sorts, ["Pcpo.cpo"]) |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
117 |
(Tactic.rtac (typedef_cpo OF cpo_thms) 1) |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
118 |
|> Theory.add_path name |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
119 |
|> PureThy.add_thms |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
120 |
([(("adm_" ^ name, admissible'), []), |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
121 |
(("cont_" ^ Rep_name, cont_Rep OF cpo_thms), []), |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
122 |
(("cont_" ^ Abs_name, cont_Abs OF cpo_thms), []), |
17833 | 123 |
(("lub_" ^ name, typedef_lub OF cpo_thms), []), |
124 |
(("thelub_" ^ name, typedef_thelub OF cpo_thms), []), |
|
125 |
(("compact_" ^ name, typedef_compact OF cpo_thms), [])]) |
|
16919
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
126 |
|>> Theory.parent_path; |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
127 |
in (theory', defs) end; |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
128 |
|
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
129 |
fun make_pcpo UUmem (theory, defs as (type_def, less_def, set_def)) = |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
130 |
let |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
131 |
val UUmem' = option_fold_rule set_def UUmem; |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
132 |
val pcpo_thms = [type_def, less_def, UUmem']; |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
133 |
val (theory', _) = |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
134 |
theory |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
135 |
|> AxClass.add_inst_arity_i (full_tname, lhs_sorts, ["Pcpo.pcpo"]) |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
136 |
(Tactic.rtac (typedef_pcpo OF pcpo_thms) 1) |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
137 |
|> Theory.add_path name |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
138 |
|> PureThy.add_thms |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
139 |
([((Rep_name ^ "_strict", Rep_strict OF pcpo_thms), []), |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
140 |
((Abs_name ^ "_strict", Abs_strict OF pcpo_thms), []), |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
141 |
((Rep_name ^ "_defined", Rep_defined OF pcpo_thms), []), |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
142 |
((Abs_name ^ "_defined", Abs_defined OF pcpo_thms), []) |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
143 |
]) |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
144 |
|>> Theory.parent_path; |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
145 |
in (theory', defs) end; |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
146 |
|
16696 | 147 |
fun pcpodef_result (theory, UUmem_admissible) = |
148 |
let |
|
149 |
val UUmem = UUmem_admissible RS conjunct1; |
|
150 |
val admissible = UUmem_admissible RS conjunct2; |
|
151 |
in |
|
152 |
theory |
|
153 |
|> make_po (Tactic.rtac exI 1 THEN Tactic.rtac UUmem 1) |
|
16919
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
154 |
|> make_cpo admissible |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
155 |
|> make_pcpo UUmem |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
156 |
|> (fn (theory', (type_def, _, _)) => (theory', type_def)) |
16696 | 157 |
end; |
158 |
||
159 |
fun cpodef_result (theory, nonempty_admissible) = |
|
160 |
let |
|
161 |
val nonempty = nonempty_admissible RS conjunct1; |
|
162 |
val admissible = nonempty_admissible RS conjunct2; |
|
163 |
in |
|
164 |
theory |
|
165 |
|> make_po (Tactic.rtac nonempty 1) |
|
16919
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
166 |
|> make_cpo admissible |
6df23e3180fb
removed duplicated code; generate new lub and thelub lemmas for new cpo types
huffman
parents:
16696
diff
changeset
|
167 |
|> (fn (theory', (type_def, _, _)) => (theory', type_def)) |
16696 | 168 |
end; |
169 |
||
170 |
in (goal, if pcpo then pcpodef_result else cpodef_result) end |
|
171 |
handle ERROR => err_in_cpodef name; |
|
172 |
||
173 |
(* cpodef_proof interface *) |
|
174 |
||
17336 | 175 |
fun gen_pcpodef_proof prep_term pcpo ((def, name), typ, set, opt_morphs) thy = |
16696 | 176 |
let |
177 |
val (goal, att) = |
|
178 |
gen_prepare_pcpodef prep_term pcpo def name typ set opt_morphs thy; |
|
17336 | 179 |
in IsarThy.theorem_i Drule.internalK ("", [att]) (goal, ([], [])) thy end; |
16696 | 180 |
|
181 |
val pcpodef_proof = gen_pcpodef_proof read_term true; |
|
182 |
val pcpodef_proof_i = gen_pcpodef_proof cert_term true; |
|
183 |
||
184 |
val cpodef_proof = gen_pcpodef_proof read_term false; |
|
185 |
val cpodef_proof_i = gen_pcpodef_proof cert_term false; |
|
186 |
||
187 |
||
188 |
(** outer syntax **) |
|
189 |
||
17057 | 190 |
local structure P = OuterParse and K = OuterKeyword in |
16696 | 191 |
|
192 |
(* copied from HOL/Tools/typedef_package.ML *) |
|
193 |
val typedef_proof_decl = |
|
194 |
Scan.optional (P.$$$ "(" |-- |
|
195 |
((P.$$$ "open" >> K false) -- Scan.option P.name || P.name >> (fn s => (true, SOME s))) |
|
196 |
--| P.$$$ ")") (true, NONE) -- |
|
197 |
(P.type_args -- P.name) -- P.opt_infix -- (P.$$$ "=" |-- P.term) -- |
|
198 |
Scan.option (P.$$$ "morphisms" |-- P.!!! (P.name -- P.name)); |
|
199 |
||
200 |
fun mk_pcpodef_proof pcpo ((((((def, opt_name), (vs, t)), mx), A), morphs)) = |
|
201 |
(if pcpo then pcpodef_proof else cpodef_proof) |
|
202 |
((def, getOpt (opt_name, Syntax.type_name t mx)), (t, vs, mx), A, morphs); |
|
203 |
||
204 |
val pcpodefP = |
|
205 |
OuterSyntax.command "pcpodef" "HOLCF type definition (requires admissibility proof)" K.thy_goal |
|
206 |
(typedef_proof_decl >> |
|
207 |
(Toplevel.print oo (Toplevel.theory_to_proof o mk_pcpodef_proof true))); |
|
208 |
||
209 |
val cpodefP = |
|
210 |
OuterSyntax.command "cpodef" "HOLCF type definition (requires admissibility proof)" K.thy_goal |
|
211 |
(typedef_proof_decl >> |
|
212 |
(Toplevel.print oo (Toplevel.theory_to_proof o mk_pcpodef_proof false))); |
|
213 |
||
214 |
val _ = OuterSyntax.add_parsers [pcpodefP, cpodefP]; |
|
215 |
||
216 |
end; |
|
217 |
||
218 |
end; |