src/HOLCF/Tools/pcpodef_package.ML
author wenzelm
Thu, 11 Dec 2008 16:50:18 +0100
changeset 29063 7619f0561cd7
parent 29060 d7bde0b4bf72
child 29585 c23295521af5
permissions -rw-r--r--
pcpodef package: state two goals, instead of encoded conjunction;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOLCF/Tools/pcpodef_package.ML
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     2
    Author:     Brian Huffman
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     3
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     4
Primitive domain definitions for HOLCF, similar to Gordon/HOL-style
29060
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
     5
typedef (see also ~~/src/HOL/Tools/typedef_package.ML).
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     6
*)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     7
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     8
signature PCPODEF_PACKAGE =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     9
sig
29060
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
    10
  val pcpodef_proof: (bool * string) * (bstring * string list * mixfix) * term
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    11
    * (string * string) option -> theory -> Proof.state
29060
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
    12
  val pcpodef_proof_cmd: (bool * string) * (bstring * string list * mixfix) * string
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    13
    * (string * string) option -> theory -> Proof.state
29060
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
    14
  val cpodef_proof: (bool * string) * (bstring * string list * mixfix) * term
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    15
    * (string * string) option -> theory -> Proof.state
29060
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
    16
  val cpodef_proof_cmd: (bool * string) * (bstring * string list * mixfix) * string
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    17
    * (string * string) option -> theory -> Proof.state
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    18
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    19
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    20
structure PcpodefPackage: PCPODEF_PACKAGE =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    21
struct
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    22
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    23
(** type definitions **)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    24
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    25
(* prepare_cpodef *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    26
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    27
fun err_in_cpodef msg name =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    28
  cat_error msg ("The error(s) above occurred in cpodef " ^ quote name);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    29
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    30
fun declare_type_name a = Variable.declare_constraints (Logic.mk_type (TFree (a, dummyS)));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    31
29060
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
    32
fun adm_const T = Const (@{const_name adm}, (T --> HOLogic.boolT) --> HOLogic.boolT);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    33
fun mk_adm (x, T, P) = adm_const T $ absfree (x, T, P);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    34
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    35
fun prepare_pcpodef prep_term pcpo def name (t, vs, mx) raw_set opt_morphs thy =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    36
  let
29060
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
    37
    val _ = Theory.requires thy "Pcpodef" "pcpodefs";
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    38
    val ctxt = ProofContext.init thy;
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28662
diff changeset
    39
    val full = Sign.full_bname thy;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    40
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    41
    (*rhs*)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    42
    val full_name = full name;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    43
    val set = prep_term (ctxt |> fold declare_type_name vs) raw_set;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    44
    val setT = Term.fastype_of set;
29060
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
    45
    val rhs_tfrees = Term.add_tfrees set [];
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    46
    val oldT = HOLogic.dest_setT setT handle TYPE _ =>
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 24867
diff changeset
    47
      error ("Not a set type: " ^ quote (Syntax.string_of_typ ctxt setT));
29060
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
    48
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
    49
    (*goal*)
29063
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 29060
diff changeset
    50
    val goal_UU_mem = HOLogic.mk_Trueprop (HOLogic.mk_mem (Const (@{const_name UU}, oldT), set));
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 29060
diff changeset
    51
    val goal_nonempty =
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 29060
diff changeset
    52
      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: 29060
diff changeset
    53
    val goal_admissible =
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 29060
diff changeset
    54
      HOLogic.mk_Trueprop (mk_adm ("x", oldT, HOLogic.mk_mem (Free ("x", oldT), set)));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    55
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    56
    (*lhs*)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    57
    val defS = Sign.defaultS thy;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    58
    val lhs_tfrees = map (fn v => (v, the_default defS (AList.lookup (op =) rhs_tfrees v))) vs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    59
    val lhs_sorts = map snd lhs_tfrees;
29060
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
    60
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    61
    val tname = Syntax.type_name t mx;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    62
    val full_tname = full tname;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    63
    val newT = Type (full_tname, map TFree lhs_tfrees);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    64
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    65
    val (Rep_name, Abs_name) = the_default ("Rep_" ^ name, "Abs_" ^ name) opt_morphs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    66
    val RepC = Const (full Rep_name, newT --> oldT);
29060
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
    67
    fun lessC T = Const (@{const_name sq_le}, T --> T --> HOLogic.boolT);
28073
5e9f00f4f209 adapted to class instantiation compliance
haftmann
parents: 27691
diff changeset
    68
    val less_def = Logic.mk_equals (lessC newT,
5e9f00f4f209 adapted to class instantiation compliance
haftmann
parents: 27691
diff changeset
    69
      Abs ("x", newT, Abs ("y", newT, lessC oldT $ (RepC $ Bound 1) $ (RepC $ Bound 0))));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    70
28073
5e9f00f4f209 adapted to class instantiation compliance
haftmann
parents: 27691
diff changeset
    71
    fun make_po tac thy1 =
5e9f00f4f209 adapted to class instantiation compliance
haftmann
parents: 27691
diff changeset
    72
      let
5e9f00f4f209 adapted to class instantiation compliance
haftmann
parents: 27691
diff changeset
    73
        val ((_, {type_definition, set_def, ...}), thy2) = thy1
28662
64ab5bb68d4c tuned typedef interface
haftmann
parents: 28394
diff changeset
    74
          |> TypedefPackage.add_typedef def (SOME name) (t, vs, mx) set opt_morphs tac;
28073
5e9f00f4f209 adapted to class instantiation compliance
haftmann
parents: 27691
diff changeset
    75
        val lthy3 = thy2
29060
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
    76
          |> TheoryTarget.instantiation ([full_tname], lhs_tfrees, @{sort po});
28073
5e9f00f4f209 adapted to class instantiation compliance
haftmann
parents: 27691
diff changeset
    77
        val less_def' = Syntax.check_term lthy3 less_def;
5e9f00f4f209 adapted to class instantiation compliance
haftmann
parents: 27691
diff changeset
    78
        val ((_, (_, less_definition')), lthy4) = lthy3
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28662
diff changeset
    79
          |> Specification.definition (NONE, ((Binding.name ("less_" ^ name ^ "_def"), []), less_def'));
28073
5e9f00f4f209 adapted to class instantiation compliance
haftmann
parents: 27691
diff changeset
    80
        val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy4);
5e9f00f4f209 adapted to class instantiation compliance
haftmann
parents: 27691
diff changeset
    81
        val less_definition = singleton (ProofContext.export lthy4 ctxt_thy) less_definition';
5e9f00f4f209 adapted to class instantiation compliance
haftmann
parents: 27691
diff changeset
    82
        val thy5 = lthy4
5e9f00f4f209 adapted to class instantiation compliance
haftmann
parents: 27691
diff changeset
    83
          |> Class.prove_instantiation_instance
29060
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
    84
              (K (Tactic.rtac (@{thm typedef_po} OF [type_definition, less_definition]) 1))
28394
b9c8e3a12a98 LocalTheory.exit_global;
wenzelm
parents: 28377
diff changeset
    85
          |> LocalTheory.exit_global;
28073
5e9f00f4f209 adapted to class instantiation compliance
haftmann
parents: 27691
diff changeset
    86
      in ((type_definition, less_definition, set_def), thy5) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    87
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    88
    fun make_cpo admissible (type_def, less_def, set_def) theory =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    89
      let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    90
        val admissible' = fold_rule (the_list set_def) admissible;
28377
73b380ba1743 proper transfer of theorems that involve classes being instantiated here;
wenzelm
parents: 28083
diff changeset
    91
        val cpo_thms = map (Thm.transfer theory) [type_def, less_def, admissible'];
73b380ba1743 proper transfer of theorems that involve classes being instantiated here;
wenzelm
parents: 28083
diff changeset
    92
        val theory' = theory
29060
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
    93
          |> AxClass.prove_arity (full_tname, lhs_sorts, @{sort cpo})
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
    94
            (Tactic.rtac (@{thm typedef_cpo} OF cpo_thms) 1);
28377
73b380ba1743 proper transfer of theorems that involve classes being instantiated here;
wenzelm
parents: 28083
diff changeset
    95
        val cpo_thms' = map (Thm.transfer theory') cpo_thms;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    96
      in
28377
73b380ba1743 proper transfer of theorems that involve classes being instantiated here;
wenzelm
parents: 28083
diff changeset
    97
        theory'
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24509
diff changeset
    98
        |> Sign.add_path name
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    99
        |> PureThy.add_thms
29060
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   100
          ([(("adm_" ^ name, admissible'), []),
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   101
            (("cont_" ^ Rep_name, @{thm typedef_cont_Rep} OF cpo_thms'), []),
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   102
            (("cont_" ^ Abs_name, @{thm typedef_cont_Abs} OF cpo_thms'), []),
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   103
            (("lub_" ^ name, @{thm typedef_lub} OF cpo_thms'), []),
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   104
            (("thelub_" ^ name, @{thm typedef_thelub} OF cpo_thms'), []),
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   105
            (("compact_" ^ name, @{thm typedef_compact} OF cpo_thms'), [])])
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   106
        |> snd
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24509
diff changeset
   107
        |> Sign.parent_path
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   108
      end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   109
29063
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 29060
diff changeset
   110
    fun make_pcpo UU_mem (type_def, less_def, set_def) theory =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   111
      let
29063
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 29060
diff changeset
   112
        val UU_mem' = fold_rule (the_list set_def) UU_mem;
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 29060
diff changeset
   113
        val pcpo_thms = map (Thm.transfer theory) [type_def, less_def, UU_mem'];
28377
73b380ba1743 proper transfer of theorems that involve classes being instantiated here;
wenzelm
parents: 28083
diff changeset
   114
        val theory' = theory
29060
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   115
          |> AxClass.prove_arity (full_tname, lhs_sorts, @{sort pcpo})
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   116
            (Tactic.rtac (@{thm typedef_pcpo} OF pcpo_thms) 1);
28377
73b380ba1743 proper transfer of theorems that involve classes being instantiated here;
wenzelm
parents: 28083
diff changeset
   117
        val pcpo_thms' = map (Thm.transfer theory') pcpo_thms;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   118
      in
28377
73b380ba1743 proper transfer of theorems that involve classes being instantiated here;
wenzelm
parents: 28083
diff changeset
   119
        theory'
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24509
diff changeset
   120
        |> Sign.add_path name
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   121
        |> PureThy.add_thms
29060
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   122
            ([((Rep_name ^ "_strict", @{thm typedef_Rep_strict} OF pcpo_thms'), []),
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   123
              ((Abs_name ^ "_strict", @{thm typedef_Abs_strict} OF pcpo_thms'), []),
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   124
              ((Rep_name ^ "_strict_iff", @{thm typedef_Rep_strict_iff} OF pcpo_thms'), []),
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   125
              ((Abs_name ^ "_strict_iff", @{thm typedef_Abs_strict_iff} OF pcpo_thms'), []),
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   126
              ((Rep_name ^ "_defined", @{thm typedef_Rep_defined} OF pcpo_thms'), []),
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   127
              ((Abs_name ^ "_defined", @{thm typedef_Abs_defined} OF pcpo_thms'), [])
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   128
              ])
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   129
        |> snd
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24509
diff changeset
   130
        |> Sign.parent_path
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   131
      end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   132
29063
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 29060
diff changeset
   133
    fun pcpodef_result UU_mem admissible =
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 29060
diff changeset
   134
      make_po (Tactic.rtac exI 1 THEN Tactic.rtac UU_mem 1)
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 29060
diff changeset
   135
      #-> (fn defs => make_cpo admissible defs #> make_pcpo UU_mem defs);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   136
29063
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 29060
diff changeset
   137
    fun cpodef_result nonempty admissible =
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 29060
diff changeset
   138
      make_po (Tactic.rtac nonempty 1)
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 29060
diff changeset
   139
      #-> make_cpo admissible;
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 29060
diff changeset
   140
  in
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 29060
diff changeset
   141
    if pcpo
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 29060
diff changeset
   142
    then (goal_UU_mem, goal_admissible, pcpodef_result)
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 29060
diff changeset
   143
    else (goal_nonempty, goal_admissible, cpodef_result)
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 29060
diff changeset
   144
  end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   145
  handle ERROR msg => err_in_cpodef msg name;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   146
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   147
29060
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   148
(* proof interface *)
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   149
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   150
local
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   151
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   152
fun gen_pcpodef_proof prep_term pcpo ((def, name), typ, set, opt_morphs) thy =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   153
  let
29063
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 29060
diff changeset
   154
    val (goal1, goal2, make_result) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   155
      prepare_pcpodef prep_term pcpo def name typ set opt_morphs thy;
29063
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 29060
diff changeset
   156
    fun after_qed [[th1, th2]] = ProofContext.theory (make_result th1 th2);
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 29060
diff changeset
   157
  in Proof.theorem_i NONE after_qed [[(goal1, []), (goal2, [])]] (ProofContext.init thy) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   158
29060
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   159
in
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   160
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   161
fun pcpodef_proof x = gen_pcpodef_proof Syntax.check_term true x;
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   162
fun pcpodef_proof_cmd x = gen_pcpodef_proof Syntax.read_term true x;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   163
29060
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   164
fun cpodef_proof x = gen_pcpodef_proof Syntax.check_term false x;
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   165
fun cpodef_proof_cmd x = gen_pcpodef_proof Syntax.read_term false x;
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   166
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   167
end;
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   168
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   169
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   170
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   171
(** outer syntax **)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   172
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   173
local structure P = OuterParse and K = OuterKeyword in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   174
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   175
val typedef_proof_decl =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   176
  Scan.optional (P.$$$ "(" |--
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   177
      ((P.$$$ "open" >> K false) -- Scan.option P.name || P.name >> (fn s => (true, SOME s)))
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   178
        --| P.$$$ ")") (true, NONE) --
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   179
    (P.type_args -- P.name) -- P.opt_infix -- (P.$$$ "=" |-- P.term) --
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   180
    Scan.option (P.$$$ "morphisms" |-- P.!!! (P.name -- P.name));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   181
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   182
fun mk_pcpodef_proof pcpo ((((((def, opt_name), (vs, t)), mx), A), morphs)) =
29060
d7bde0b4bf72 tuned signature: pcpodef_proof, pcpodef_proof_cmd etc.;
wenzelm
parents: 28965
diff changeset
   183
  (if pcpo then pcpodef_proof_cmd else cpodef_proof_cmd)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   184
    ((def, the_default (Syntax.type_name t mx) opt_name), (t, vs, mx), A, morphs);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   185
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24712
diff changeset
   186
val _ =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   187
  OuterSyntax.command "pcpodef" "HOLCF type definition (requires admissibility proof)" K.thy_goal
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   188
    (typedef_proof_decl >>
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   189
      (Toplevel.print oo (Toplevel.theory_to_proof o mk_pcpodef_proof true)));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   190
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24712
diff changeset
   191
val _ =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   192
  OuterSyntax.command "cpodef" "HOLCF type definition (requires admissibility proof)" K.thy_goal
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   193
    (typedef_proof_decl >>
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   194
      (Toplevel.print oo (Toplevel.theory_to_proof o mk_pcpodef_proof false)));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   195
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   196
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   197
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   198
end;