| author | wenzelm | 
| Mon, 19 Dec 2022 11:16:46 +0100 | |
| changeset 76701 | 3543ecb4c97d | 
| parent 74305 | 28a582aa25dd | 
| child 80634 | a90ab1ea6458 | 
| permissions | -rw-r--r-- | 
| 42151 | 1  | 
(* Title: HOL/HOLCF/Tools/cpodef.ML  | 
| 23152 | 2  | 
Author: Brian Huffman  | 
3  | 
||
4  | 
Primitive domain definitions for HOLCF, similar to Gordon/HOL-style  | 
|
| 
31738
 
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
 
haftmann 
parents: 
31725 
diff
changeset
 | 
5  | 
typedef (see also ~~/src/HOL/Tools/typedef.ML).  | 
| 23152 | 6  | 
*)  | 
7  | 
||
| 40772 | 8  | 
signature CPODEF =  | 
| 23152 | 9  | 
sig  | 
| 33646 | 10  | 
type cpo_info =  | 
11  | 
    { below_def: thm, adm: thm, cont_Rep: thm, cont_Abs: thm,
 | 
|
| 
41028
 
0acff85f95c7
cpodef no longer generates lemma is_lub_foo, since it is redundant with lub_foo
 
huffman 
parents: 
40889 
diff
changeset
 | 
12  | 
lub: thm, compact: thm }  | 
| 33646 | 13  | 
type pcpo_info =  | 
| 
41029
 
f7d8cfa6e7fc
pcpodef no longer generates _defined lemmas, use _bottom_iff lemmas instead
 
huffman 
parents: 
41028 
diff
changeset
 | 
14  | 
    { Rep_strict: thm, Abs_strict: thm,
 | 
| 
 
f7d8cfa6e7fc
pcpodef no longer generates _defined lemmas, use _bottom_iff lemmas instead
 
huffman 
parents: 
41028 
diff
changeset
 | 
15  | 
Rep_bottom_iff: thm, Abs_bottom_iff: thm }  | 
| 33646 | 16  | 
|
| 
49759
 
ecf1d5d87e5e
removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
 
huffman 
parents: 
46961 
diff
changeset
 | 
17  | 
val add_podef: binding * (string * sort) list * mixfix ->  | 
| 61110 | 18  | 
term -> Typedef.bindings option -> (Proof.context -> tactic) -> theory ->  | 
| 33646 | 19  | 
(Typedef.info * thm) * theory  | 
| 
49759
 
ecf1d5d87e5e
removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
 
huffman 
parents: 
46961 
diff
changeset
 | 
20  | 
val add_cpodef: binding * (string * sort) list * mixfix ->  | 
| 61110 | 21  | 
term -> Typedef.bindings option -> (Proof.context -> tactic) * (Proof.context -> tactic) ->  | 
| 58959 | 22  | 
theory -> (Typedef.info * cpo_info) * theory  | 
| 
49759
 
ecf1d5d87e5e
removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
 
huffman 
parents: 
46961 
diff
changeset
 | 
23  | 
val add_pcpodef: binding * (string * sort) list * mixfix ->  | 
| 61110 | 24  | 
term -> Typedef.bindings option -> (Proof.context -> tactic) * (Proof.context -> tactic) ->  | 
| 58959 | 25  | 
theory -> (Typedef.info * cpo_info * pcpo_info) * theory  | 
| 33646 | 26  | 
|
| 
49759
 
ecf1d5d87e5e
removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
 
huffman 
parents: 
46961 
diff
changeset
 | 
27  | 
val cpodef_proof:  | 
| 
 
ecf1d5d87e5e
removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
 
huffman 
parents: 
46961 
diff
changeset
 | 
28  | 
(binding * (string * sort) list * mixfix) * term  | 
| 61110 | 29  | 
* Typedef.bindings option -> theory -> Proof.state  | 
| 
49759
 
ecf1d5d87e5e
removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
 
huffman 
parents: 
46961 
diff
changeset
 | 
30  | 
val cpodef_proof_cmd:  | 
| 
 
ecf1d5d87e5e
removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
 
huffman 
parents: 
46961 
diff
changeset
 | 
31  | 
(binding * (string * string option) list * mixfix) * string  | 
| 61110 | 32  | 
* Typedef.bindings option -> theory -> Proof.state  | 
| 
49759
 
ecf1d5d87e5e
removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
 
huffman 
parents: 
46961 
diff
changeset
 | 
33  | 
val pcpodef_proof:  | 
| 
 
ecf1d5d87e5e
removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
 
huffman 
parents: 
46961 
diff
changeset
 | 
34  | 
(binding * (string * sort) list * mixfix) * term  | 
| 61110 | 35  | 
* Typedef.bindings option -> theory -> Proof.state  | 
| 
49759
 
ecf1d5d87e5e
removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
 
huffman 
parents: 
46961 
diff
changeset
 | 
36  | 
val pcpodef_proof_cmd:  | 
| 
 
ecf1d5d87e5e
removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
 
huffman 
parents: 
46961 
diff
changeset
 | 
37  | 
(binding * (string * string option) list * mixfix) * string  | 
| 61110 | 38  | 
* Typedef.bindings option -> theory -> Proof.state  | 
| 40832 | 39  | 
end  | 
| 23152 | 40  | 
|
| 
41296
 
6aaf80ea9715
switch to transparent ascription, to avoid warning messages
 
huffman 
parents: 
41029 
diff
changeset
 | 
41  | 
structure Cpodef : CPODEF =  | 
| 23152 | 42  | 
struct  | 
43  | 
||
44  | 
(** type definitions **)  | 
|
45  | 
||
| 33646 | 46  | 
type cpo_info =  | 
47  | 
  { below_def: thm, adm: thm, cont_Rep: thm, cont_Abs: thm,
 | 
|
| 
41028
 
0acff85f95c7
cpodef no longer generates lemma is_lub_foo, since it is redundant with lub_foo
 
huffman 
parents: 
40889 
diff
changeset
 | 
48  | 
lub: thm, compact: thm }  | 
| 23152 | 49  | 
|
| 33646 | 50  | 
type pcpo_info =  | 
| 
41029
 
f7d8cfa6e7fc
pcpodef no longer generates _defined lemmas, use _bottom_iff lemmas instead
 
huffman 
parents: 
41028 
diff
changeset
 | 
51  | 
  { Rep_strict: thm, Abs_strict: thm,
 | 
| 
 
f7d8cfa6e7fc
pcpodef no longer generates _defined lemmas, use _bottom_iff lemmas instead
 
huffman 
parents: 
41028 
diff
changeset
 | 
52  | 
Rep_bottom_iff: thm, Abs_bottom_iff: thm }  | 
| 33646 | 53  | 
|
54  | 
(* building terms *)  | 
|
| 23152 | 55  | 
|
| 74305 | 56  | 
fun adm_const T = \<^Const>\<open>adm T\<close>  | 
| 44241 | 57  | 
fun mk_adm (x, T, P) = adm_const T $ absfree (x, T) P  | 
| 23152 | 58  | 
|
| 74305 | 59  | 
fun below_const T = \<^Const>\<open>below T\<close>  | 
| 33646 | 60  | 
|
61  | 
(* proving class instances *)  | 
|
62  | 
||
63  | 
fun prove_cpo  | 
|
64  | 
(name: binding)  | 
|
65  | 
(newT: typ)  | 
|
| 61110 | 66  | 
opt_bindings  | 
| 33646 | 67  | 
(type_definition: thm) (* type_definition Rep Abs A *)  | 
68  | 
(below_def: thm) (* op << == %x y. Rep x << Rep y *)  | 
|
69  | 
(admissible: thm) (* adm (%x. x : set) *)  | 
|
70  | 
(thy: theory)  | 
|
71  | 
=  | 
|
72  | 
let  | 
|
| 61110 | 73  | 
    val {Rep_name, Abs_name, ...} = Typedef.make_bindings name opt_bindings;
 | 
| 49833 | 74  | 
val cpo_thms = map (Thm.transfer thy) [type_definition, below_def, admissible]  | 
| 40832 | 75  | 
val (full_tname, Ts) = dest_Type newT  | 
76  | 
val lhs_sorts = map (snd o dest_TFree) Ts  | 
|
| 60754 | 77  | 
    fun tac ctxt = resolve_tac ctxt [@{thm typedef_cpo} OF cpo_thms] 1
 | 
| 69597 | 78  | 
val thy = Axclass.prove_arity (full_tname, lhs_sorts, \<^sort>\<open>cpo\<close>) tac thy  | 
| 
33678
 
2a2014cbb2a6
cleaned up, removed unneeded call to Syntax.check_term
 
huffman 
parents: 
33646 
diff
changeset
 | 
79  | 
(* transfer thms so that they will know about the new cpo instance *)  | 
| 40832 | 80  | 
val cpo_thms' = map (Thm.transfer thy) cpo_thms  | 
81  | 
fun make thm = Drule.zero_var_indexes (thm OF cpo_thms')  | 
|
82  | 
    val cont_Rep = make @{thm typedef_cont_Rep}
 | 
|
83  | 
    val cont_Abs = make @{thm typedef_cont_Abs}
 | 
|
84  | 
    val lub = make @{thm typedef_lub}
 | 
|
85  | 
    val compact = make @{thm typedef_compact}
 | 
|
| 35902 | 86  | 
val (_, thy) =  | 
87  | 
thy  | 
|
| 33646 | 88  | 
|> Sign.add_path (Binding.name_of name)  | 
| 
39557
 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 
wenzelm 
parents: 
38756 
diff
changeset
 | 
89  | 
|> Global_Theory.add_thms  | 
| 49833 | 90  | 
([((Binding.prefix_name "adm_" name, admissible), []),  | 
91  | 
((Binding.prefix_name "cont_" Rep_name, cont_Rep ), []),  | 
|
92  | 
((Binding.prefix_name "cont_" Abs_name, cont_Abs ), []),  | 
|
93  | 
((Binding.prefix_name "lub_" name, lub ), []),  | 
|
94  | 
((Binding.prefix_name "compact_" name, compact ), [])])  | 
|
| 40832 | 95  | 
||> Sign.parent_path  | 
| 33646 | 96  | 
val cpo_info : cpo_info =  | 
| 49833 | 97  | 
      { below_def = below_def, adm = admissible, cont_Rep = cont_Rep,
 | 
| 
41028
 
0acff85f95c7
cpodef no longer generates lemma is_lub_foo, since it is redundant with lub_foo
 
huffman 
parents: 
40889 
diff
changeset
 | 
98  | 
cont_Abs = cont_Abs, lub = lub, compact = compact }  | 
| 33646 | 99  | 
in  | 
| 35902 | 100  | 
(cpo_info, thy)  | 
| 40832 | 101  | 
end  | 
| 33646 | 102  | 
|
103  | 
fun prove_pcpo  | 
|
104  | 
(name: binding)  | 
|
105  | 
(newT: typ)  | 
|
| 61110 | 106  | 
opt_bindings  | 
| 33646 | 107  | 
(type_definition: thm) (* type_definition Rep Abs A *)  | 
108  | 
(below_def: thm) (* op << == %x y. Rep x << Rep y *)  | 
|
| 
41429
 
cf5f025bc3c7
renamed constant 'UU' to 'bottom', keeping 'UU' as alternative input syntax;
 
huffman 
parents: 
41296 
diff
changeset
 | 
109  | 
(bottom_mem: thm) (* bottom : set *)  | 
| 33646 | 110  | 
(thy: theory)  | 
111  | 
=  | 
|
112  | 
let  | 
|
| 61110 | 113  | 
    val {Rep_name, Abs_name, ...} = Typedef.make_bindings name opt_bindings;
 | 
| 49833 | 114  | 
val pcpo_thms = map (Thm.transfer thy) [type_definition, below_def, bottom_mem]  | 
| 40832 | 115  | 
val (full_tname, Ts) = dest_Type newT  | 
116  | 
val lhs_sorts = map (snd o dest_TFree) Ts  | 
|
| 60754 | 117  | 
    fun tac ctxt = resolve_tac ctxt [@{thm typedef_pcpo} OF pcpo_thms] 1
 | 
| 69597 | 118  | 
val thy = Axclass.prove_arity (full_tname, lhs_sorts, \<^sort>\<open>pcpo\<close>) tac thy  | 
| 40832 | 119  | 
val pcpo_thms' = map (Thm.transfer thy) pcpo_thms  | 
120  | 
fun make thm = Drule.zero_var_indexes (thm OF pcpo_thms')  | 
|
121  | 
    val Rep_strict = make @{thm typedef_Rep_strict}
 | 
|
122  | 
    val Abs_strict = make @{thm typedef_Abs_strict}
 | 
|
123  | 
    val Rep_bottom_iff = make @{thm typedef_Rep_bottom_iff}
 | 
|
124  | 
    val Abs_bottom_iff = make @{thm typedef_Abs_bottom_iff}
 | 
|
| 35902 | 125  | 
val (_, thy) =  | 
126  | 
thy  | 
|
| 33646 | 127  | 
|> Sign.add_path (Binding.name_of name)  | 
| 
39557
 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 
wenzelm 
parents: 
38756 
diff
changeset
 | 
128  | 
|> Global_Theory.add_thms  | 
| 35902 | 129  | 
([((Binding.suffix_name "_strict" Rep_name, Rep_strict), []),  | 
130  | 
((Binding.suffix_name "_strict" Abs_name, Abs_strict), []),  | 
|
| 
40321
 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 
huffman 
parents: 
39557 
diff
changeset
 | 
131  | 
((Binding.suffix_name "_bottom_iff" Rep_name, Rep_bottom_iff), []),  | 
| 
41029
 
f7d8cfa6e7fc
pcpodef no longer generates _defined lemmas, use _bottom_iff lemmas instead
 
huffman 
parents: 
41028 
diff
changeset
 | 
132  | 
((Binding.suffix_name "_bottom_iff" Abs_name, Abs_bottom_iff), [])])  | 
| 40832 | 133  | 
||> Sign.parent_path  | 
| 33646 | 134  | 
val pcpo_info =  | 
135  | 
      { Rep_strict = Rep_strict, Abs_strict = Abs_strict,
 | 
|
| 
41029
 
f7d8cfa6e7fc
pcpodef no longer generates _defined lemmas, use _bottom_iff lemmas instead
 
huffman 
parents: 
41028 
diff
changeset
 | 
136  | 
Rep_bottom_iff = Rep_bottom_iff, Abs_bottom_iff = Abs_bottom_iff }  | 
| 33646 | 137  | 
in  | 
| 35902 | 138  | 
(pcpo_info, thy)  | 
| 40832 | 139  | 
end  | 
| 33646 | 140  | 
|
141  | 
(* prepare_cpodef *)  | 
|
142  | 
||
| 61110 | 143  | 
fun prepare prep_term name (tname, raw_args, _) raw_set thy =  | 
| 23152 | 144  | 
let  | 
145  | 
(*rhs*)  | 
|
| 
36153
 
1ac501e16a6a
replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
 
wenzelm 
parents: 
35994 
diff
changeset
 | 
146  | 
val tmp_ctxt =  | 
| 42361 | 147  | 
Proof_Context.init_global thy  | 
| 40832 | 148  | 
|> fold (Variable.declare_typ o TFree) raw_args  | 
149  | 
val set = prep_term tmp_ctxt raw_set  | 
|
150  | 
val tmp_ctxt' = tmp_ctxt |> Variable.declare_term set  | 
|
| 
35840
 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 
wenzelm 
parents: 
35742 
diff
changeset
 | 
151  | 
|
| 40832 | 152  | 
val setT = Term.fastype_of set  | 
| 23152 | 153  | 
val oldT = HOLogic.dest_setT setT handle TYPE _ =>  | 
| 40832 | 154  | 
      error ("Not a set type: " ^ quote (Syntax.string_of_typ tmp_ctxt setT))
 | 
| 
29060
 
d7bde0b4bf72
tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
 
wenzelm 
parents: 
28965 
diff
changeset
 | 
155  | 
|
| 33646 | 156  | 
(*lhs*)  | 
| 42361 | 157  | 
val lhs_tfrees = map (Proof_Context.check_tfree tmp_ctxt') raw_args  | 
| 40832 | 158  | 
val full_tname = Sign.full_name thy tname  | 
159  | 
val newT = Type (full_tname, map TFree lhs_tfrees)  | 
|
| 33646 | 160  | 
in  | 
| 61110 | 161  | 
(newT, oldT, set)  | 
| 33646 | 162  | 
end  | 
163  | 
||
| 61110 | 164  | 
fun add_podef typ set opt_bindings tac thy =  | 
| 33646 | 165  | 
let  | 
| 
49759
 
ecf1d5d87e5e
removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
 
huffman 
parents: 
46961 
diff
changeset
 | 
166  | 
val name = #1 typ  | 
| 69829 | 167  | 
    val ((full_tname, info as ({Rep_name, ...}, {type_definition, ...})), thy) =
 | 
168  | 
thy  | 
|
169  | 
|> Named_Target.theory_map_result (apsnd o Typedef.transform_info)  | 
|
170  | 
           (Typedef.add_typedef {overloaded = false} typ set opt_bindings tac)
 | 
|
| 40832 | 171  | 
val oldT = #rep_type (#1 info)  | 
172  | 
val newT = #abs_type (#1 info)  | 
|
173  | 
val lhs_tfrees = map dest_TFree (snd (dest_Type newT))  | 
|
| 33646 | 174  | 
|
| 40832 | 175  | 
val RepC = Const (Rep_name, newT --> oldT)  | 
| 
33678
 
2a2014cbb2a6
cleaned up, removed unneeded call to Syntax.check_term
 
huffman 
parents: 
33646 
diff
changeset
 | 
176  | 
val below_eqn = Logic.mk_equals (below_const newT,  | 
| 40832 | 177  | 
      Abs ("x", newT, Abs ("y", newT, below_const oldT $ (RepC $ Bound 1) $ (RepC $ Bound 0))))
 | 
| 40889 | 178  | 
val ((_, (_, below_ldef)), lthy) = thy  | 
| 69597 | 179  | 
|> Class.instantiation ([full_tname], lhs_tfrees, \<^sort>\<open>po\<close>)  | 
| 63180 | 180  | 
|> Specification.definition NONE [] []  | 
| 
63064
 
2f18172214c8
support 'assumes' in specifications, e.g. 'definition', 'inductive';
 
wenzelm 
parents: 
61260 
diff
changeset
 | 
181  | 
((Binding.prefix_name "below_" (Thm.def_binding name), []), below_eqn)  | 
| 42361 | 182  | 
val ctxt_thy = Proof_Context.init_global (Proof_Context.theory_of lthy)  | 
183  | 
val below_def = singleton (Proof_Context.export lthy ctxt_thy) below_ldef  | 
|
| 40889 | 184  | 
val thy = lthy  | 
185  | 
|> Class.prove_instantiation_exit  | 
|
| 60754 | 186  | 
          (fn ctxt => resolve_tac ctxt [@{thm typedef_po} OF [type_definition, below_def]] 1)
 | 
| 40889 | 187  | 
in ((info, below_def), thy) end  | 
| 33646 | 188  | 
|
189  | 
fun prepare_cpodef  | 
|
190  | 
(prep_term: Proof.context -> 'a -> term)  | 
|
| 
35840
 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 
wenzelm 
parents: 
35742 
diff
changeset
 | 
191  | 
(typ: binding * (string * sort) list * mixfix)  | 
| 33646 | 192  | 
(raw_set: 'a)  | 
| 61110 | 193  | 
opt_bindings  | 
| 33646 | 194  | 
(thy: theory)  | 
195  | 
: term * term * (thm -> thm -> theory -> (Typedef.info * cpo_info) * theory) =  | 
|
196  | 
let  | 
|
| 
49759
 
ecf1d5d87e5e
removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
 
huffman 
parents: 
46961 
diff
changeset
 | 
197  | 
val name = #1 typ  | 
| 61110 | 198  | 
val (newT, oldT, set) = prepare prep_term name typ raw_set thy  | 
| 33646 | 199  | 
|
| 
29063
 
7619f0561cd7
pcpodef package: state two goals, instead of encoded conjunction;
 
wenzelm 
parents: 
29060 
diff
changeset
 | 
200  | 
val goal_nonempty =  | 
| 40832 | 201  | 
      HOLogic.mk_Trueprop (HOLogic.mk_exists ("x", oldT, HOLogic.mk_mem (Free ("x", oldT), set)))
 | 
| 
29063
 
7619f0561cd7
pcpodef package: state two goals, instead of encoded conjunction;
 
wenzelm 
parents: 
29060 
diff
changeset
 | 
202  | 
val goal_admissible =  | 
| 40832 | 203  | 
      HOLogic.mk_Trueprop (mk_adm ("x", oldT, HOLogic.mk_mem (Free ("x", oldT), set)))
 | 
| 23152 | 204  | 
|
| 33646 | 205  | 
fun cpodef_result nonempty admissible thy =  | 
| 28073 | 206  | 
let  | 
| 49833 | 207  | 
        val ((info as (_, {type_definition, ...}), below_def), thy) = thy
 | 
| 61110 | 208  | 
|> add_podef typ set opt_bindings (fn ctxt => resolve_tac ctxt [nonempty] 1)  | 
| 40889 | 209  | 
val (cpo_info, thy) = thy  | 
| 61110 | 210  | 
|> prove_cpo name newT opt_bindings type_definition below_def admissible  | 
| 23152 | 211  | 
in  | 
| 40889 | 212  | 
((info, cpo_info), thy)  | 
| 40832 | 213  | 
end  | 
| 
29063
 
7619f0561cd7
pcpodef package: state two goals, instead of encoded conjunction;
 
wenzelm 
parents: 
29060 
diff
changeset
 | 
214  | 
in  | 
| 33646 | 215  | 
(goal_nonempty, goal_admissible, cpodef_result)  | 
| 
29063
 
7619f0561cd7
pcpodef package: state two goals, instead of encoded conjunction;
 
wenzelm 
parents: 
29060 
diff
changeset
 | 
216  | 
end  | 
| 30345 | 217  | 
handle ERROR msg =>  | 
| 
49759
 
ecf1d5d87e5e
removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
 
huffman 
parents: 
46961 
diff
changeset
 | 
218  | 
    cat_error msg ("The error(s) above occurred in cpodef " ^ Binding.print (#1 typ))
 | 
| 23152 | 219  | 
|
| 33646 | 220  | 
fun prepare_pcpodef  | 
221  | 
(prep_term: Proof.context -> 'a -> term)  | 
|
| 
35840
 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 
wenzelm 
parents: 
35742 
diff
changeset
 | 
222  | 
(typ: binding * (string * sort) list * mixfix)  | 
| 33646 | 223  | 
(raw_set: 'a)  | 
| 61110 | 224  | 
opt_bindings  | 
| 33646 | 225  | 
(thy: theory)  | 
226  | 
: term * term * (thm -> thm -> theory -> (Typedef.info * cpo_info * pcpo_info) * theory) =  | 
|
227  | 
let  | 
|
| 
49759
 
ecf1d5d87e5e
removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
 
huffman 
parents: 
46961 
diff
changeset
 | 
228  | 
val name = #1 typ  | 
| 61110 | 229  | 
val (newT, oldT, set) = prepare prep_term name typ raw_set thy  | 
| 33646 | 230  | 
|
| 
41429
 
cf5f025bc3c7
renamed constant 'UU' to 'bottom', keeping 'UU' as alternative input syntax;
 
huffman 
parents: 
41296 
diff
changeset
 | 
231  | 
val goal_bottom_mem =  | 
| 74305 | 232  | 
HOLogic.mk_Trueprop (HOLogic.mk_mem (\<^Const>\<open>bottom oldT\<close>, set))  | 
| 33646 | 233  | 
|
234  | 
val goal_admissible =  | 
|
| 40832 | 235  | 
      HOLogic.mk_Trueprop (mk_adm ("x", oldT, HOLogic.mk_mem (Free ("x", oldT), set)))
 | 
| 33646 | 236  | 
|
| 
41429
 
cf5f025bc3c7
renamed constant 'UU' to 'bottom', keeping 'UU' as alternative input syntax;
 
huffman 
parents: 
41296 
diff
changeset
 | 
237  | 
fun pcpodef_result bottom_mem admissible thy =  | 
| 33646 | 238  | 
let  | 
| 60754 | 239  | 
fun tac ctxt = resolve_tac ctxt [exI] 1 THEN resolve_tac ctxt [bottom_mem] 1  | 
| 49833 | 240  | 
        val ((info as (_, {type_definition, ...}), below_def), thy) = thy
 | 
| 61110 | 241  | 
|> add_podef typ set opt_bindings tac  | 
| 40889 | 242  | 
val (cpo_info, thy) = thy  | 
| 61110 | 243  | 
|> prove_cpo name newT opt_bindings type_definition below_def admissible  | 
| 40889 | 244  | 
val (pcpo_info, thy) = thy  | 
| 61110 | 245  | 
|> prove_pcpo name newT opt_bindings type_definition below_def bottom_mem  | 
| 33646 | 246  | 
in  | 
| 40889 | 247  | 
((info, cpo_info, pcpo_info), thy)  | 
| 40832 | 248  | 
end  | 
| 33646 | 249  | 
in  | 
| 
41429
 
cf5f025bc3c7
renamed constant 'UU' to 'bottom', keeping 'UU' as alternative input syntax;
 
huffman 
parents: 
41296 
diff
changeset
 | 
250  | 
(goal_bottom_mem, goal_admissible, pcpodef_result)  | 
| 33646 | 251  | 
end  | 
252  | 
handle ERROR msg =>  | 
|
| 
49759
 
ecf1d5d87e5e
removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
 
huffman 
parents: 
46961 
diff
changeset
 | 
253  | 
    cat_error msg ("The error(s) above occurred in pcpodef " ^ Binding.print (#1 typ))
 | 
| 33646 | 254  | 
|
| 23152 | 255  | 
|
| 
33645
 
562635ab559b
use Drule.standard (following typedef package), add pcpodef tactic interface
 
huffman 
parents: 
33553 
diff
changeset
 | 
256  | 
(* tactic interface *)  | 
| 
 
562635ab559b
use Drule.standard (following typedef package), add pcpodef tactic interface
 
huffman 
parents: 
33553 
diff
changeset
 | 
257  | 
|
| 61110 | 258  | 
fun add_cpodef typ set opt_bindings (tac1, tac2) thy =  | 
| 33646 | 259  | 
let  | 
260  | 
val (goal1, goal2, cpodef_result) =  | 
|
| 61110 | 261  | 
prepare_cpodef Syntax.check_term typ set opt_bindings thy  | 
| 58959 | 262  | 
val thm1 = Goal.prove_global thy [] [] goal1 (tac1 o #context)  | 
| 33646 | 263  | 
handle ERROR msg => cat_error msg  | 
| 40832 | 264  | 
        ("Failed to prove non-emptiness of " ^ quote (Syntax.string_of_term_global thy set))
 | 
| 58959 | 265  | 
val thm2 = Goal.prove_global thy [] [] goal2 (tac2 o #context)  | 
| 33646 | 266  | 
handle ERROR msg => cat_error msg  | 
| 40832 | 267  | 
        ("Failed to prove admissibility of " ^ quote (Syntax.string_of_term_global thy set))
 | 
268  | 
in cpodef_result thm1 thm2 thy end  | 
|
| 
33645
 
562635ab559b
use Drule.standard (following typedef package), add pcpodef tactic interface
 
huffman 
parents: 
33553 
diff
changeset
 | 
269  | 
|
| 61110 | 270  | 
fun add_pcpodef typ set opt_bindings (tac1, tac2) thy =  | 
| 
33645
 
562635ab559b
use Drule.standard (following typedef package), add pcpodef tactic interface
 
huffman 
parents: 
33553 
diff
changeset
 | 
271  | 
let  | 
| 
 
562635ab559b
use Drule.standard (following typedef package), add pcpodef tactic interface
 
huffman 
parents: 
33553 
diff
changeset
 | 
272  | 
val (goal1, goal2, pcpodef_result) =  | 
| 61110 | 273  | 
prepare_pcpodef Syntax.check_term typ set opt_bindings thy  | 
| 58959 | 274  | 
val thm1 = Goal.prove_global thy [] [] goal1 (tac1 o #context)  | 
| 
33645
 
562635ab559b
use Drule.standard (following typedef package), add pcpodef tactic interface
 
huffman 
parents: 
33553 
diff
changeset
 | 
275  | 
handle ERROR msg => cat_error msg  | 
| 40832 | 276  | 
        ("Failed to prove non-emptiness of " ^ quote (Syntax.string_of_term_global thy set))
 | 
| 58959 | 277  | 
val thm2 = Goal.prove_global thy [] [] goal2 (tac2 o #context)  | 
| 
33645
 
562635ab559b
use Drule.standard (following typedef package), add pcpodef tactic interface
 
huffman 
parents: 
33553 
diff
changeset
 | 
278  | 
handle ERROR msg => cat_error msg  | 
| 40832 | 279  | 
        ("Failed to prove admissibility of " ^ quote (Syntax.string_of_term_global thy set))
 | 
280  | 
in pcpodef_result thm1 thm2 thy end  | 
|
| 
33645
 
562635ab559b
use Drule.standard (following typedef package), add pcpodef tactic interface
 
huffman 
parents: 
33553 
diff
changeset
 | 
281  | 
|
| 
 
562635ab559b
use Drule.standard (following typedef package), add pcpodef tactic interface
 
huffman 
parents: 
33553 
diff
changeset
 | 
282  | 
|
| 
29060
 
d7bde0b4bf72
tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
 
wenzelm 
parents: 
28965 
diff
changeset
 | 
283  | 
(* proof interface *)  | 
| 
 
d7bde0b4bf72
tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
 
wenzelm 
parents: 
28965 
diff
changeset
 | 
284  | 
|
| 
 
d7bde0b4bf72
tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
 
wenzelm 
parents: 
28965 
diff
changeset
 | 
285  | 
local  | 
| 23152 | 286  | 
|
| 
35840
 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 
wenzelm 
parents: 
35742 
diff
changeset
 | 
287  | 
fun gen_cpodef_proof prep_term prep_constraint  | 
| 61110 | 288  | 
((b, raw_args, mx), set, opt_bindings) thy =  | 
| 23152 | 289  | 
let  | 
| 42361 | 290  | 
val ctxt = Proof_Context.init_global thy  | 
| 40832 | 291  | 
val args = map (apsnd (prep_constraint ctxt)) raw_args  | 
| 
29063
 
7619f0561cd7
pcpodef package: state two goals, instead of encoded conjunction;
 
wenzelm 
parents: 
29060 
diff
changeset
 | 
292  | 
val (goal1, goal2, make_result) =  | 
| 61110 | 293  | 
prepare_cpodef prep_term (b, args, mx) set opt_bindings thy  | 
| 42361 | 294  | 
fun after_qed [[th1, th2]] = Proof_Context.background_theory (snd o make_result th1 th2)  | 
| 40832 | 295  | 
| after_qed _ = raise Fail "cpodef_proof"  | 
296  | 
in Proof.theorem NONE after_qed [[(goal1, []), (goal2, [])]] ctxt end  | 
|
| 33646 | 297  | 
|
| 
35840
 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 
wenzelm 
parents: 
35742 
diff
changeset
 | 
298  | 
fun gen_pcpodef_proof prep_term prep_constraint  | 
| 61110 | 299  | 
((b, raw_args, mx), set, opt_bindings) thy =  | 
| 33646 | 300  | 
let  | 
| 42361 | 301  | 
val ctxt = Proof_Context.init_global thy  | 
| 40832 | 302  | 
val args = map (apsnd (prep_constraint ctxt)) raw_args  | 
| 33646 | 303  | 
val (goal1, goal2, make_result) =  | 
| 61110 | 304  | 
prepare_pcpodef prep_term (b, args, mx) set opt_bindings thy  | 
| 42361 | 305  | 
fun after_qed [[th1, th2]] = Proof_Context.background_theory (snd o make_result th1 th2)  | 
| 40832 | 306  | 
| after_qed _ = raise Fail "pcpodef_proof"  | 
307  | 
in Proof.theorem NONE after_qed [[(goal1, []), (goal2, [])]] ctxt end  | 
|
| 23152 | 308  | 
|
| 
29060
 
d7bde0b4bf72
tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
 
wenzelm 
parents: 
28965 
diff
changeset
 | 
309  | 
in  | 
| 
 
d7bde0b4bf72
tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
 
wenzelm 
parents: 
28965 
diff
changeset
 | 
310  | 
|
| 40832 | 311  | 
fun cpodef_proof x = gen_cpodef_proof Syntax.check_term (K I) x  | 
312  | 
fun cpodef_proof_cmd x = gen_cpodef_proof Syntax.read_term Typedecl.read_constraint x  | 
|
| 23152 | 313  | 
|
| 40832 | 314  | 
fun pcpodef_proof x = gen_pcpodef_proof Syntax.check_term (K I) x  | 
315  | 
fun pcpodef_proof_cmd x = gen_pcpodef_proof Syntax.read_term Typedecl.read_constraint x  | 
|
| 
29060
 
d7bde0b4bf72
tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
 
wenzelm 
parents: 
28965 
diff
changeset
 | 
316  | 
|
| 40832 | 317  | 
end  | 
| 
29060
 
d7bde0b4bf72
tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
 
wenzelm 
parents: 
28965 
diff
changeset
 | 
318  | 
|
| 23152 | 319  | 
|
320  | 
||
321  | 
(** outer syntax **)  | 
|
322  | 
||
| 61110 | 323  | 
local  | 
324  | 
||
325  | 
fun cpodef pcpo =  | 
|
| 
49759
 
ecf1d5d87e5e
removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
 
huffman 
parents: 
46961 
diff
changeset
 | 
326  | 
(Parse.type_args_constrained -- Parse.binding) -- Parse.opt_mixfix --  | 
| 69597 | 327  | 
(\<^keyword>\<open>=\<close> |-- Parse.term) --  | 
| 
49759
 
ecf1d5d87e5e
removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
 
huffman 
parents: 
46961 
diff
changeset
 | 
328  | 
Scan.option  | 
| 69597 | 329  | 
(\<^keyword>\<open>morphisms\<close> |-- Parse.!!! (Parse.binding -- Parse.binding))  | 
| 61110 | 330  | 
>> (fn ((((args, t), mx), A), morphs) =>  | 
331  | 
Toplevel.theory_to_proof  | 
|
332  | 
((if pcpo then pcpodef_proof_cmd else cpodef_proof_cmd)  | 
|
333  | 
((t, args, mx), A, SOME (Typedef.make_morphisms t morphs))))  | 
|
| 23152 | 334  | 
|
| 61110 | 335  | 
in  | 
| 23152 | 336  | 
|
| 24867 | 337  | 
val _ =  | 
| 69597 | 338  | 
Outer_Syntax.command \<^command_keyword>\<open>pcpodef\<close>  | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46949 
diff
changeset
 | 
339  | 
"HOLCF type definition (requires admissibility proof)"  | 
| 61110 | 340  | 
(cpodef true)  | 
| 23152 | 341  | 
|
| 24867 | 342  | 
val _ =  | 
| 69597 | 343  | 
Outer_Syntax.command \<^command_keyword>\<open>cpodef\<close>  | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46949 
diff
changeset
 | 
344  | 
"HOLCF type definition (requires admissibility proof)"  | 
| 61110 | 345  | 
(cpodef false)  | 
| 23152 | 346  | 
|
| 40832 | 347  | 
end  | 
| 61110 | 348  | 
|
349  | 
end  |