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