| author | blanchet | 
| Fri, 03 Sep 2010 13:45:12 +0200 | |
| changeset 39111 | 2e9bdc6fbedf | 
| parent 38756 | d07959fabde6 | 
| child 39557 | fe5722fce758 | 
| permissions | -rw-r--r-- | 
| 31738 
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31725diff
changeset | 1 | (* Title: HOLCF/Tools/pcpodef.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: 
31725diff
changeset | 5 | typedef (see also ~~/src/HOL/Tools/typedef.ML). | 
| 23152 | 6 | *) | 
| 7 | ||
| 31738 
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31725diff
changeset | 8 | signature PCPODEF = | 
| 23152 | 9 | sig | 
| 33646 | 10 | type cpo_info = | 
| 11 |     { below_def: thm, adm: thm, cont_Rep: thm, cont_Abs: thm,
 | |
| 12 | lub: thm, thelub: thm, compact: thm } | |
| 13 | type pcpo_info = | |
| 14 |     { Rep_strict: thm, Abs_strict: thm, Rep_strict_iff: thm, Abs_strict_iff: thm,
 | |
| 15 | Rep_defined: thm, Abs_defined: thm } | |
| 16 | ||
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 17 | val add_podef: bool -> binding option -> binding * (string * sort) list * mixfix -> | 
| 33646 | 18 | term -> (binding * binding) option -> tactic -> theory -> | 
| 19 | (Typedef.info * thm) * theory | |
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 20 | val add_cpodef: bool -> binding option -> binding * (string * sort) list * mixfix -> | 
| 33646 | 21 | term -> (binding * binding) option -> tactic * tactic -> theory -> | 
| 22 | (Typedef.info * cpo_info) * theory | |
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 23 | val add_pcpodef: bool -> binding option -> binding * (string * sort) list * mixfix -> | 
| 33646 | 24 | term -> (binding * binding) option -> tactic * tactic -> theory -> | 
| 25 | (Typedef.info * cpo_info * pcpo_info) * theory | |
| 26 | ||
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 27 | val cpodef_proof: (bool * binding) | 
| 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 28 | * (binding * (string * sort) list * mixfix) * term | 
| 33646 | 29 | * (binding * binding) option -> theory -> Proof.state | 
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 30 | val cpodef_proof_cmd: (bool * binding) | 
| 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 31 | * (binding * (string * string option) list * mixfix) * string | 
| 33646 | 32 | * (binding * binding) option -> theory -> Proof.state | 
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 33 | val pcpodef_proof: (bool * binding) | 
| 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 34 | * (binding * (string * sort) list * mixfix) * term | 
| 30345 | 35 | * (binding * binding) option -> theory -> Proof.state | 
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 36 | val pcpodef_proof_cmd: (bool * binding) | 
| 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 37 | * (binding * (string * string option) list * mixfix) * string | 
| 30345 | 38 | * (binding * binding) option -> theory -> Proof.state | 
| 23152 | 39 | end; | 
| 40 | ||
| 31738 
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31725diff
changeset | 41 | structure Pcpodef :> PCPODEF = | 
| 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,
 | |
| 48 | lub: thm, thelub: thm, compact: thm } | |
| 23152 | 49 | |
| 33646 | 50 | type pcpo_info = | 
| 51 |   { Rep_strict: thm, Abs_strict: thm, Rep_strict_iff: thm, Abs_strict_iff: thm,
 | |
| 52 | Rep_defined: thm, Abs_defined: thm } | |
| 53 | ||
| 54 | (* building terms *) | |
| 23152 | 55 | |
| 29060 
d7bde0b4bf72
tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
 wenzelm parents: 
28965diff
changeset | 56 | fun adm_const T = Const (@{const_name adm}, (T --> HOLogic.boolT) --> HOLogic.boolT);
 | 
| 23152 | 57 | fun mk_adm (x, T, P) = adm_const T $ absfree (x, T, P); | 
| 58 | ||
| 33646 | 59 | fun below_const T = Const (@{const_name below}, T --> T --> HOLogic.boolT);
 | 
| 60 | ||
| 61 | (* manipulating theorems *) | |
| 62 | ||
| 63 | fun fold_adm_mem thm NONE = thm | |
| 64 | | fold_adm_mem thm (SOME set_def) = | |
| 65 |     let val rule = @{lemma "A == B ==> adm (%x. x : B) ==> adm (%x. x : A)" by simp}
 | |
| 66 | in rule OF [set_def, thm] end; | |
| 67 | ||
| 68 | fun fold_UU_mem thm NONE = thm | |
| 69 | | fold_UU_mem thm (SOME set_def) = | |
| 70 |     let val rule = @{lemma "A == B ==> UU : B ==> UU : A" by simp}
 | |
| 71 | in rule OF [set_def, thm] end; | |
| 72 | ||
| 73 | (* proving class instances *) | |
| 74 | ||
| 75 | fun prove_cpo | |
| 76 | (name: binding) | |
| 77 | (newT: typ) | |
| 78 | (Rep_name: binding, Abs_name: binding) | |
| 79 | (type_definition: thm) (* type_definition Rep Abs A *) | |
| 80 | (set_def: thm option) (* A == set *) | |
| 81 | (below_def: thm) (* op << == %x y. Rep x << Rep y *) | |
| 82 | (admissible: thm) (* adm (%x. x : set) *) | |
| 83 | (thy: theory) | |
| 84 | = | |
| 85 | let | |
| 86 | val admissible' = fold_adm_mem admissible set_def; | |
| 87 | val cpo_thms = map (Thm.transfer thy) [type_definition, below_def, admissible']; | |
| 88 | val (full_tname, Ts) = dest_Type newT; | |
| 89 | val lhs_sorts = map (snd o dest_TFree) Ts; | |
| 35902 | 90 |     val tac = Tactic.rtac (@{thm typedef_cpo} OF cpo_thms) 1;
 | 
| 91 |     val thy = AxClass.prove_arity (full_tname, lhs_sorts, @{sort cpo}) tac thy;
 | |
| 33678 
2a2014cbb2a6
cleaned up, removed unneeded call to Syntax.check_term
 huffman parents: 
33646diff
changeset | 92 | (* transfer thms so that they will know about the new cpo instance *) | 
| 35902 | 93 | val cpo_thms' = map (Thm.transfer thy) cpo_thms; | 
| 36241 
2a4cec6bcae2
replace many uses of Drule.export_without_context with Drule.zero_var_indexes
 huffman parents: 
36153diff
changeset | 94 | fun make thm = Drule.zero_var_indexes (thm OF cpo_thms'); | 
| 35902 | 95 |     val cont_Rep = make @{thm typedef_cont_Rep};
 | 
| 96 |     val cont_Abs = make @{thm typedef_cont_Abs};
 | |
| 97 |     val lub = make @{thm typedef_lub};
 | |
| 98 |     val thelub = make @{thm typedef_thelub};
 | |
| 99 |     val compact = make @{thm typedef_compact};
 | |
| 100 | val (_, thy) = | |
| 101 | thy | |
| 33646 | 102 | |> Sign.add_path (Binding.name_of name) | 
| 103 | |> PureThy.add_thms | |
| 35902 | 104 | ([((Binding.prefix_name "adm_" name, admissible'), []), | 
| 105 | ((Binding.prefix_name "cont_" Rep_name, cont_Rep ), []), | |
| 106 | ((Binding.prefix_name "cont_" Abs_name, cont_Abs ), []), | |
| 107 | ((Binding.prefix_name "lub_" name, lub ), []), | |
| 108 | ((Binding.prefix_name "thelub_" name, thelub ), []), | |
| 109 | ((Binding.prefix_name "compact_" name, compact ), [])]) | |
| 110 | ||> Sign.parent_path; | |
| 33646 | 111 | val cpo_info : cpo_info = | 
| 112 |       { below_def = below_def, adm = admissible', cont_Rep = cont_Rep,
 | |
| 113 | cont_Abs = cont_Abs, lub = lub, thelub = thelub, compact = compact }; | |
| 114 | in | |
| 35902 | 115 | (cpo_info, thy) | 
| 33646 | 116 | end; | 
| 117 | ||
| 118 | fun prove_pcpo | |
| 119 | (name: binding) | |
| 120 | (newT: typ) | |
| 121 | (Rep_name: binding, Abs_name: binding) | |
| 122 | (type_definition: thm) (* type_definition Rep Abs A *) | |
| 123 | (set_def: thm option) (* A == set *) | |
| 124 | (below_def: thm) (* op << == %x y. Rep x << Rep y *) | |
| 125 | (UU_mem: thm) (* UU : set *) | |
| 126 | (thy: theory) | |
| 127 | = | |
| 128 | let | |
| 129 | val UU_mem' = fold_UU_mem UU_mem set_def; | |
| 130 | val pcpo_thms = map (Thm.transfer thy) [type_definition, below_def, UU_mem']; | |
| 131 | val (full_tname, Ts) = dest_Type newT; | |
| 132 | val lhs_sorts = map (snd o dest_TFree) Ts; | |
| 35902 | 133 |     val tac = Tactic.rtac (@{thm typedef_pcpo} OF pcpo_thms) 1;
 | 
| 134 |     val thy = AxClass.prove_arity (full_tname, lhs_sorts, @{sort pcpo}) tac thy;
 | |
| 135 | val pcpo_thms' = map (Thm.transfer thy) pcpo_thms; | |
| 36241 
2a4cec6bcae2
replace many uses of Drule.export_without_context with Drule.zero_var_indexes
 huffman parents: 
36153diff
changeset | 136 | fun make thm = Drule.zero_var_indexes (thm OF pcpo_thms'); | 
| 35902 | 137 |     val Rep_strict = make @{thm typedef_Rep_strict};
 | 
| 138 |     val Abs_strict = make @{thm typedef_Abs_strict};
 | |
| 139 |     val Rep_strict_iff = make @{thm typedef_Rep_strict_iff};
 | |
| 140 |     val Abs_strict_iff = make @{thm typedef_Abs_strict_iff};
 | |
| 141 |     val Rep_defined = make @{thm typedef_Rep_defined};
 | |
| 142 |     val Abs_defined = make @{thm typedef_Abs_defined};
 | |
| 143 | val (_, thy) = | |
| 144 | thy | |
| 33646 | 145 | |> Sign.add_path (Binding.name_of name) | 
| 146 | |> PureThy.add_thms | |
| 35902 | 147 | ([((Binding.suffix_name "_strict" Rep_name, Rep_strict), []), | 
| 148 | ((Binding.suffix_name "_strict" Abs_name, Abs_strict), []), | |
| 149 | ((Binding.suffix_name "_strict_iff" Rep_name, Rep_strict_iff), []), | |
| 150 | ((Binding.suffix_name "_strict_iff" Abs_name, Abs_strict_iff), []), | |
| 151 | ((Binding.suffix_name "_defined" Rep_name, Rep_defined), []), | |
| 152 | ((Binding.suffix_name "_defined" Abs_name, Abs_defined), [])]) | |
| 153 | ||> Sign.parent_path; | |
| 33646 | 154 | val pcpo_info = | 
| 155 |       { Rep_strict = Rep_strict, Abs_strict = Abs_strict,
 | |
| 156 | Rep_strict_iff = Rep_strict_iff, Abs_strict_iff = Abs_strict_iff, | |
| 157 | Rep_defined = Rep_defined, Abs_defined = Abs_defined }; | |
| 158 | in | |
| 35902 | 159 | (pcpo_info, thy) | 
| 33646 | 160 | end; | 
| 161 | ||
| 162 | (* prepare_cpodef *) | |
| 163 | ||
| 164 | fun declare_type_name a = | |
| 165 | Variable.declare_constraints (Logic.mk_type (TFree (a, dummyS))); | |
| 166 | ||
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 167 | fun prepare prep_term name (tname, raw_args, mx) raw_set opt_morphs thy = | 
| 23152 | 168 | let | 
| 29060 
d7bde0b4bf72
tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
 wenzelm parents: 
28965diff
changeset | 169 | val _ = Theory.requires thy "Pcpodef" "pcpodefs"; | 
| 30345 | 170 | |
| 23152 | 171 | (*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: 
35994diff
changeset | 172 | val tmp_ctxt = | 
| 36610 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 wenzelm parents: 
36323diff
changeset | 173 | ProofContext.init_global thy | 
| 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: 
35994diff
changeset | 174 | |> fold (Variable.declare_typ o TFree) raw_args; | 
| 
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: 
35994diff
changeset | 175 | val set = prep_term tmp_ctxt raw_set; | 
| 
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: 
35994diff
changeset | 176 | val tmp_ctxt' = tmp_ctxt |> Variable.declare_term set; | 
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 177 | |
| 23152 | 178 | val setT = Term.fastype_of set; | 
| 179 | val oldT = HOLogic.dest_setT setT handle TYPE _ => | |
| 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: 
35994diff
changeset | 180 |       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: 
28965diff
changeset | 181 | |
| 33646 | 182 | (*lhs*) | 
| 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: 
35994diff
changeset | 183 | val lhs_tfrees = map (ProofContext.check_tfree tmp_ctxt') raw_args; | 
| 33646 | 184 | val full_tname = Sign.full_name thy tname; | 
| 185 | val newT = Type (full_tname, map TFree lhs_tfrees); | |
| 186 | ||
| 187 | val morphs = opt_morphs | |
| 188 | |> the_default (Binding.prefix_name "Rep_" name, Binding.prefix_name "Abs_" name); | |
| 189 | in | |
| 33678 
2a2014cbb2a6
cleaned up, removed unneeded call to Syntax.check_term
 huffman parents: 
33646diff
changeset | 190 | (newT, oldT, set, morphs) | 
| 33646 | 191 | end | 
| 192 | ||
| 193 | fun add_podef def opt_name typ set opt_morphs tac thy = | |
| 194 | let | |
| 195 | val name = the_default (#1 typ) opt_name; | |
| 35994 
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
 wenzelm parents: 
35912diff
changeset | 196 |     val ((full_tname, info as ({Rep_name, ...}, {type_definition, set_def, ...})), thy2) = thy
 | 
| 35742 | 197 | |> Typedef.add_typedef_global def opt_name typ set opt_morphs tac; | 
| 35994 
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
 wenzelm parents: 
35912diff
changeset | 198 | val oldT = #rep_type (#1 info); | 
| 
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
 wenzelm parents: 
35912diff
changeset | 199 | val newT = #abs_type (#1 info); | 
| 33646 | 200 | val lhs_tfrees = map dest_TFree (snd (dest_Type newT)); | 
| 201 | ||
| 202 | val RepC = Const (Rep_name, newT --> oldT); | |
| 33678 
2a2014cbb2a6
cleaned up, removed unneeded call to Syntax.check_term
 huffman parents: 
33646diff
changeset | 203 | val below_eqn = Logic.mk_equals (below_const newT, | 
| 33646 | 204 |       Abs ("x", newT, Abs ("y", newT, below_const oldT $ (RepC $ Bound 1) $ (RepC $ Bound 0))));
 | 
| 205 | val lthy3 = thy2 | |
| 38348 
cf7b2121ad9d
moved instantiation target formally to class_target.ML
 haftmann parents: 
36960diff
changeset | 206 |       |> Class.instantiation ([full_tname], lhs_tfrees, @{sort po});
 | 
| 33678 
2a2014cbb2a6
cleaned up, removed unneeded call to Syntax.check_term
 huffman parents: 
33646diff
changeset | 207 | val ((_, (_, below_ldef)), lthy4) = lthy3 | 
| 33646 | 208 | |> Specification.definition (NONE, | 
| 33678 
2a2014cbb2a6
cleaned up, removed unneeded call to Syntax.check_term
 huffman parents: 
33646diff
changeset | 209 | ((Binding.prefix_name "below_" (Binding.suffix_name "_def" name), []), below_eqn)); | 
| 36610 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 wenzelm parents: 
36323diff
changeset | 210 | val ctxt_thy = ProofContext.init_global (ProofContext.theory_of lthy4); | 
| 33678 
2a2014cbb2a6
cleaned up, removed unneeded call to Syntax.check_term
 huffman parents: 
33646diff
changeset | 211 | val below_def = singleton (ProofContext.export lthy4 ctxt_thy) below_ldef; | 
| 33646 | 212 | val thy5 = lthy4 | 
| 213 | |> Class.prove_instantiation_instance | |
| 33678 
2a2014cbb2a6
cleaned up, removed unneeded call to Syntax.check_term
 huffman parents: 
33646diff
changeset | 214 |           (K (Tactic.rtac (@{thm typedef_po} OF [type_definition, below_def]) 1))
 | 
| 33671 | 215 | |> Local_Theory.exit_global; | 
| 33678 
2a2014cbb2a6
cleaned up, removed unneeded call to Syntax.check_term
 huffman parents: 
33646diff
changeset | 216 | in ((info, below_def), thy5) end; | 
| 33646 | 217 | |
| 218 | fun prepare_cpodef | |
| 219 | (prep_term: Proof.context -> 'a -> term) | |
| 220 | (def: bool) | |
| 221 | (name: binding) | |
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 222 | (typ: binding * (string * sort) list * mixfix) | 
| 33646 | 223 | (raw_set: 'a) | 
| 224 | (opt_morphs: (binding * binding) option) | |
| 225 | (thy: theory) | |
| 226 | : term * term * (thm -> thm -> theory -> (Typedef.info * cpo_info) * theory) = | |
| 227 | let | |
| 33678 
2a2014cbb2a6
cleaned up, removed unneeded call to Syntax.check_term
 huffman parents: 
33646diff
changeset | 228 | val (newT, oldT, set, morphs as (Rep_name, Abs_name)) = | 
| 
2a2014cbb2a6
cleaned up, removed unneeded call to Syntax.check_term
 huffman parents: 
33646diff
changeset | 229 | prepare prep_term name typ raw_set opt_morphs thy; | 
| 33646 | 230 | |
| 29063 
7619f0561cd7
pcpodef package: state two goals, instead of encoded conjunction;
 wenzelm parents: 
29060diff
changeset | 231 | val goal_nonempty = | 
| 
7619f0561cd7
pcpodef package: state two goals, instead of encoded conjunction;
 wenzelm parents: 
29060diff
changeset | 232 |       HOLogic.mk_Trueprop (HOLogic.mk_exists ("x", oldT, HOLogic.mk_mem (Free ("x", oldT), set)));
 | 
| 
7619f0561cd7
pcpodef package: state two goals, instead of encoded conjunction;
 wenzelm parents: 
29060diff
changeset | 233 | val goal_admissible = | 
| 
7619f0561cd7
pcpodef package: state two goals, instead of encoded conjunction;
 wenzelm parents: 
29060diff
changeset | 234 |       HOLogic.mk_Trueprop (mk_adm ("x", oldT, HOLogic.mk_mem (Free ("x", oldT), set)));
 | 
| 23152 | 235 | |
| 33646 | 236 | fun cpodef_result nonempty admissible thy = | 
| 28073 | 237 | let | 
| 35994 
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
 wenzelm parents: 
35912diff
changeset | 238 |         val ((info as (_, {type_definition, set_def, ...}), below_def), thy2) = thy
 | 
| 33646 | 239 | |> add_podef def (SOME name) typ set opt_morphs (Tactic.rtac nonempty 1); | 
| 240 | val (cpo_info, thy3) = thy2 | |
| 241 | |> prove_cpo name newT morphs type_definition set_def below_def admissible; | |
| 23152 | 242 | in | 
| 33646 | 243 | ((info, cpo_info), thy3) | 
| 23152 | 244 | end; | 
| 29063 
7619f0561cd7
pcpodef package: state two goals, instead of encoded conjunction;
 wenzelm parents: 
29060diff
changeset | 245 | in | 
| 33646 | 246 | (goal_nonempty, goal_admissible, cpodef_result) | 
| 29063 
7619f0561cd7
pcpodef package: state two goals, instead of encoded conjunction;
 wenzelm parents: 
29060diff
changeset | 247 | end | 
| 30345 | 248 | handle ERROR msg => | 
| 249 |     cat_error msg ("The error(s) above occurred in cpodef " ^ quote (Binding.str_of name));
 | |
| 23152 | 250 | |
| 33646 | 251 | fun prepare_pcpodef | 
| 252 | (prep_term: Proof.context -> 'a -> term) | |
| 253 | (def: bool) | |
| 254 | (name: binding) | |
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 255 | (typ: binding * (string * sort) list * mixfix) | 
| 33646 | 256 | (raw_set: 'a) | 
| 257 | (opt_morphs: (binding * binding) option) | |
| 258 | (thy: theory) | |
| 259 | : term * term * (thm -> thm -> theory -> (Typedef.info * cpo_info * pcpo_info) * theory) = | |
| 260 | let | |
| 33678 
2a2014cbb2a6
cleaned up, removed unneeded call to Syntax.check_term
 huffman parents: 
33646diff
changeset | 261 | val (newT, oldT, set, morphs as (Rep_name, Abs_name)) = | 
| 
2a2014cbb2a6
cleaned up, removed unneeded call to Syntax.check_term
 huffman parents: 
33646diff
changeset | 262 | prepare prep_term name typ raw_set opt_morphs thy; | 
| 33646 | 263 | |
| 264 | val goal_UU_mem = | |
| 265 |       HOLogic.mk_Trueprop (HOLogic.mk_mem (Const (@{const_name UU}, oldT), set));
 | |
| 266 | ||
| 267 | val goal_admissible = | |
| 268 |       HOLogic.mk_Trueprop (mk_adm ("x", oldT, HOLogic.mk_mem (Free ("x", oldT), set)));
 | |
| 269 | ||
| 270 | fun pcpodef_result UU_mem admissible thy = | |
| 271 | let | |
| 272 | val tac = Tactic.rtac exI 1 THEN Tactic.rtac UU_mem 1; | |
| 35994 
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
 wenzelm parents: 
35912diff
changeset | 273 |         val ((info as (_, {type_definition, set_def, ...}), below_def), thy2) = thy
 | 
| 33646 | 274 | |> add_podef def (SOME name) typ set opt_morphs tac; | 
| 275 | val (cpo_info, thy3) = thy2 | |
| 276 | |> prove_cpo name newT morphs type_definition set_def below_def admissible; | |
| 277 | val (pcpo_info, thy4) = thy3 | |
| 278 | |> prove_pcpo name newT morphs type_definition set_def below_def UU_mem; | |
| 279 | in | |
| 280 | ((info, cpo_info, pcpo_info), thy4) | |
| 281 | end; | |
| 282 | in | |
| 283 | (goal_UU_mem, goal_admissible, pcpodef_result) | |
| 284 | end | |
| 285 | handle ERROR msg => | |
| 286 |     cat_error msg ("The error(s) above occurred in pcpodef " ^ quote (Binding.str_of name));
 | |
| 287 | ||
| 23152 | 288 | |
| 33645 
562635ab559b
use Drule.standard (following typedef package), add pcpodef tactic interface
 huffman parents: 
33553diff
changeset | 289 | (* tactic interface *) | 
| 
562635ab559b
use Drule.standard (following typedef package), add pcpodef tactic interface
 huffman parents: 
33553diff
changeset | 290 | |
| 33646 | 291 | fun add_cpodef def opt_name typ set opt_morphs (tac1, tac2) thy = | 
| 292 | let | |
| 293 | val name = the_default (#1 typ) opt_name; | |
| 294 | val (goal1, goal2, cpodef_result) = | |
| 295 | prepare_cpodef Syntax.check_term def name typ set opt_morphs thy; | |
| 296 | val thm1 = Goal.prove_global thy [] [] goal1 (K tac1) | |
| 297 | handle ERROR msg => cat_error msg | |
| 298 |         ("Failed to prove non-emptiness of " ^ quote (Syntax.string_of_term_global thy set));
 | |
| 299 | val thm2 = Goal.prove_global thy [] [] goal2 (K tac2) | |
| 300 | handle ERROR msg => cat_error msg | |
| 301 |         ("Failed to prove admissibility of " ^ quote (Syntax.string_of_term_global thy set));
 | |
| 302 | in cpodef_result thm1 thm2 thy end; | |
| 33645 
562635ab559b
use Drule.standard (following typedef package), add pcpodef tactic interface
 huffman parents: 
33553diff
changeset | 303 | |
| 33646 | 304 | fun add_pcpodef def opt_name typ set opt_morphs (tac1, tac2) thy = | 
| 33645 
562635ab559b
use Drule.standard (following typedef package), add pcpodef tactic interface
 huffman parents: 
33553diff
changeset | 305 | let | 
| 
562635ab559b
use Drule.standard (following typedef package), add pcpodef tactic interface
 huffman parents: 
33553diff
changeset | 306 | val name = the_default (#1 typ) opt_name; | 
| 
562635ab559b
use Drule.standard (following typedef package), add pcpodef tactic interface
 huffman parents: 
33553diff
changeset | 307 | val (goal1, goal2, pcpodef_result) = | 
| 33646 | 308 | prepare_pcpodef Syntax.check_term def name typ set opt_morphs thy; | 
| 309 | val thm1 = Goal.prove_global thy [] [] goal1 (K tac1) | |
| 33645 
562635ab559b
use Drule.standard (following typedef package), add pcpodef tactic interface
 huffman parents: 
33553diff
changeset | 310 | handle ERROR msg => cat_error msg | 
| 
562635ab559b
use Drule.standard (following typedef package), add pcpodef tactic interface
 huffman parents: 
33553diff
changeset | 311 |         ("Failed to prove non-emptiness of " ^ quote (Syntax.string_of_term_global thy set));
 | 
| 33646 | 312 | val thm2 = Goal.prove_global thy [] [] goal2 (K tac2) | 
| 33645 
562635ab559b
use Drule.standard (following typedef package), add pcpodef tactic interface
 huffman parents: 
33553diff
changeset | 313 | handle ERROR msg => cat_error msg | 
| 
562635ab559b
use Drule.standard (following typedef package), add pcpodef tactic interface
 huffman parents: 
33553diff
changeset | 314 |         ("Failed to prove admissibility of " ^ quote (Syntax.string_of_term_global thy set));
 | 
| 
562635ab559b
use Drule.standard (following typedef package), add pcpodef tactic interface
 huffman parents: 
33553diff
changeset | 315 | in pcpodef_result thm1 thm2 thy end; | 
| 
562635ab559b
use Drule.standard (following typedef package), add pcpodef tactic interface
 huffman parents: 
33553diff
changeset | 316 | |
| 
562635ab559b
use Drule.standard (following typedef package), add pcpodef tactic interface
 huffman parents: 
33553diff
changeset | 317 | |
| 29060 
d7bde0b4bf72
tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
 wenzelm parents: 
28965diff
changeset | 318 | (* proof interface *) | 
| 
d7bde0b4bf72
tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
 wenzelm parents: 
28965diff
changeset | 319 | |
| 
d7bde0b4bf72
tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
 wenzelm parents: 
28965diff
changeset | 320 | local | 
| 23152 | 321 | |
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 322 | fun gen_cpodef_proof prep_term prep_constraint | 
| 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 323 | ((def, name), (b, raw_args, mx), set, opt_morphs) thy = | 
| 23152 | 324 | let | 
| 36610 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 wenzelm parents: 
36323diff
changeset | 325 | val ctxt = ProofContext.init_global thy; | 
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 326 | val args = map (apsnd (prep_constraint ctxt)) raw_args; | 
| 29063 
7619f0561cd7
pcpodef package: state two goals, instead of encoded conjunction;
 wenzelm parents: 
29060diff
changeset | 327 | val (goal1, goal2, make_result) = | 
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 328 | prepare_cpodef prep_term def name (b, args, mx) set opt_morphs thy; | 
| 38756 
d07959fabde6
renamed ProofContext.theory(_result) to ProofContext.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 wenzelm parents: 
38348diff
changeset | 329 | fun after_qed [[th1, th2]] = ProofContext.background_theory (snd o make_result th1 th2) | 
| 35912 | 330 | | after_qed _ = raise Fail "cpodef_proof"; | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36241diff
changeset | 331 | in Proof.theorem NONE after_qed [[(goal1, []), (goal2, [])]] ctxt end; | 
| 33646 | 332 | |
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 333 | fun gen_pcpodef_proof prep_term prep_constraint | 
| 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 334 | ((def, name), (b, raw_args, mx), set, opt_morphs) thy = | 
| 33646 | 335 | let | 
| 36610 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 wenzelm parents: 
36323diff
changeset | 336 | val ctxt = ProofContext.init_global thy; | 
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 337 | val args = map (apsnd (prep_constraint ctxt)) raw_args; | 
| 33646 | 338 | val (goal1, goal2, make_result) = | 
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 339 | prepare_pcpodef prep_term def name (b, args, mx) set opt_morphs thy; | 
| 38756 
d07959fabde6
renamed ProofContext.theory(_result) to ProofContext.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 wenzelm parents: 
38348diff
changeset | 340 | fun after_qed [[th1, th2]] = ProofContext.background_theory (snd o make_result th1 th2) | 
| 35912 | 341 | | after_qed _ = raise Fail "pcpodef_proof"; | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36241diff
changeset | 342 | in Proof.theorem NONE after_qed [[(goal1, []), (goal2, [])]] ctxt end; | 
| 23152 | 343 | |
| 29060 
d7bde0b4bf72
tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
 wenzelm parents: 
28965diff
changeset | 344 | in | 
| 
d7bde0b4bf72
tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
 wenzelm parents: 
28965diff
changeset | 345 | |
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 346 | fun cpodef_proof x = gen_cpodef_proof Syntax.check_term (K I) x; | 
| 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 347 | fun cpodef_proof_cmd x = gen_cpodef_proof Syntax.read_term Typedecl.read_constraint x; | 
| 23152 | 348 | |
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 349 | fun pcpodef_proof x = gen_pcpodef_proof Syntax.check_term (K I) x; | 
| 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 350 | 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: 
28965diff
changeset | 351 | |
| 
d7bde0b4bf72
tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
 wenzelm parents: 
28965diff
changeset | 352 | end; | 
| 
d7bde0b4bf72
tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
 wenzelm parents: 
28965diff
changeset | 353 | |
| 23152 | 354 | |
| 355 | ||
| 356 | (** outer syntax **) | |
| 357 | ||
| 358 | val typedef_proof_decl = | |
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 359 |   Scan.optional (Parse.$$$ "(" |--
 | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 360 | ((Parse.$$$ "open" >> K false) -- Scan.option Parse.binding || | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 361 | Parse.binding >> (fn s => (true, SOME s))) | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 362 | --| Parse.$$$ ")") (true, NONE) -- | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 363 | (Parse.type_args_constrained -- Parse.binding) -- Parse.opt_mixfix -- | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 364 | (Parse.$$$ "=" |-- Parse.term) -- | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 365 | Scan.option (Parse.$$$ "morphisms" |-- Parse.!!! (Parse.binding -- Parse.binding)); | 
| 23152 | 366 | |
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 367 | fun mk_pcpodef_proof pcpo ((((((def, opt_name), (args, t)), mx), A), morphs)) = | 
| 29060 
d7bde0b4bf72
tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
 wenzelm parents: 
28965diff
changeset | 368 | (if pcpo then pcpodef_proof_cmd else cpodef_proof_cmd) | 
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35742diff
changeset | 369 | ((def, the_default t opt_name), (t, args, mx), A, morphs); | 
| 23152 | 370 | |
| 24867 | 371 | val _ = | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 372 | Outer_Syntax.command "pcpodef" "HOLCF type definition (requires admissibility proof)" | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 373 | Keyword.thy_goal | 
| 23152 | 374 | (typedef_proof_decl >> | 
| 375 | (Toplevel.print oo (Toplevel.theory_to_proof o mk_pcpodef_proof true))); | |
| 376 | ||
| 24867 | 377 | val _ = | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 378 | Outer_Syntax.command "cpodef" "HOLCF type definition (requires admissibility proof)" | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 379 | Keyword.thy_goal | 
| 23152 | 380 | (typedef_proof_decl >> | 
| 381 | (Toplevel.print oo (Toplevel.theory_to_proof o mk_pcpodef_proof false))); | |
| 382 | ||
| 383 | end; |