src/HOLCF/Tools/pcpodef_package.ML
changeset 31741 41788a3ffd6a
parent 31732 052399f580cf
parent 31740 002da20f442e
child 31743 ce6c75e7ab20
equal deleted inserted replaced
31732:052399f580cf 31741:41788a3ffd6a
     1 (*  Title:      HOLCF/Tools/pcpodef_package.ML
       
     2     Author:     Brian Huffman
       
     3 
       
     4 Primitive domain definitions for HOLCF, similar to Gordon/HOL-style
       
     5 typedef (see also ~~/src/HOL/Tools/typedef_package.ML).
       
     6 *)
       
     7 
       
     8 signature PCPODEF_PACKAGE =
       
     9 sig
       
    10   val pcpodef_proof: (bool * binding) * (binding * string list * mixfix) * term
       
    11     * (binding * binding) option -> theory -> Proof.state
       
    12   val pcpodef_proof_cmd: (bool * binding) * (binding * string list * mixfix) * string
       
    13     * (binding * binding) option -> theory -> Proof.state
       
    14   val cpodef_proof: (bool * binding) * (binding * string list * mixfix) * term
       
    15     * (binding * binding) option -> theory -> Proof.state
       
    16   val cpodef_proof_cmd: (bool * binding) * (binding * string list * mixfix) * string
       
    17     * (binding * binding) option -> theory -> Proof.state
       
    18 end;
       
    19 
       
    20 structure PcpodefPackage :> PCPODEF_PACKAGE =
       
    21 struct
       
    22 
       
    23 (** type definitions **)
       
    24 
       
    25 (* prepare_cpodef *)
       
    26 
       
    27 fun declare_type_name a = Variable.declare_constraints (Logic.mk_type (TFree (a, dummyS)));
       
    28 
       
    29 fun adm_const T = Const (@{const_name adm}, (T --> HOLogic.boolT) --> HOLogic.boolT);
       
    30 fun mk_adm (x, T, P) = adm_const T $ absfree (x, T, P);
       
    31 
       
    32 fun prepare_pcpodef prep_term pcpo def name (t, vs, mx) raw_set opt_morphs thy =
       
    33   let
       
    34     val _ = Theory.requires thy "Pcpodef" "pcpodefs";
       
    35     val ctxt = ProofContext.init thy;
       
    36 
       
    37     val full = Sign.full_name thy;
       
    38     val full_name = full name;
       
    39     val bname = Binding.name_of name;
       
    40 
       
    41     (*rhs*)
       
    42     val set = prep_term (ctxt |> fold declare_type_name vs) raw_set;
       
    43     val setT = Term.fastype_of set;
       
    44     val rhs_tfrees = Term.add_tfrees set [];
       
    45     val oldT = HOLogic.dest_setT setT handle TYPE _ =>
       
    46       error ("Not a set type: " ^ quote (Syntax.string_of_typ ctxt setT));
       
    47 
       
    48     (*goal*)
       
    49     val goal_UU_mem = HOLogic.mk_Trueprop (HOLogic.mk_mem (Const (@{const_name UU}, oldT), set));
       
    50     val goal_nonempty =
       
    51       HOLogic.mk_Trueprop (HOLogic.mk_exists ("x", oldT, HOLogic.mk_mem (Free ("x", oldT), set)));
       
    52     val goal_admissible =
       
    53       HOLogic.mk_Trueprop (mk_adm ("x", oldT, HOLogic.mk_mem (Free ("x", oldT), set)));
       
    54 
       
    55     (*lhs*)
       
    56     val defS = Sign.defaultS thy;
       
    57     val lhs_tfrees = map (fn v => (v, the_default defS (AList.lookup (op =) rhs_tfrees v))) vs;
       
    58     val lhs_sorts = map snd lhs_tfrees;
       
    59 
       
    60     val tname = Binding.map_name (Syntax.type_name mx) t;
       
    61     val full_tname = full tname;
       
    62     val newT = Type (full_tname, map TFree lhs_tfrees);
       
    63 
       
    64     val (Rep_name, Abs_name) =
       
    65       (case opt_morphs of
       
    66         NONE => (Binding.prefix_name "Rep_" name, Binding.prefix_name "Abs_" name)
       
    67       | SOME morphs => morphs);
       
    68     val RepC = Const (full Rep_name, newT --> oldT);
       
    69     fun belowC T = Const (@{const_name below}, T --> T --> HOLogic.boolT);
       
    70     val below_def = Logic.mk_equals (belowC newT,
       
    71       Abs ("x", newT, Abs ("y", newT, belowC oldT $ (RepC $ Bound 1) $ (RepC $ Bound 0))));
       
    72 
       
    73     fun make_po tac thy1 =
       
    74       let
       
    75         val ((_, {type_definition, set_def, ...}), thy2) = thy1
       
    76           |> Typedef.add_typedef def (SOME name) (t, vs, mx) set opt_morphs tac;
       
    77         val lthy3 = thy2
       
    78           |> TheoryTarget.instantiation ([full_tname], lhs_tfrees, @{sort po});
       
    79         val below_def' = Syntax.check_term lthy3 below_def;
       
    80         val ((_, (_, below_definition')), lthy4) = lthy3
       
    81           |> Specification.definition (NONE,
       
    82               ((Binding.prefix_name "below_" (Binding.suffix_name "_def" name), []), below_def'));
       
    83         val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy4);
       
    84         val below_definition = singleton (ProofContext.export lthy4 ctxt_thy) below_definition';
       
    85         val thy5 = lthy4
       
    86           |> Class.prove_instantiation_instance
       
    87               (K (Tactic.rtac (@{thm typedef_po} OF [type_definition, below_definition]) 1))
       
    88           |> LocalTheory.exit_global;
       
    89       in ((type_definition, below_definition, set_def), thy5) end;
       
    90 
       
    91     fun make_cpo admissible (type_def, below_def, set_def) theory =
       
    92       let
       
    93         val admissible' = fold_rule (the_list set_def) admissible;
       
    94         val cpo_thms = map (Thm.transfer theory) [type_def, below_def, admissible'];
       
    95         val theory' = theory
       
    96           |> AxClass.prove_arity (full_tname, lhs_sorts, @{sort cpo})
       
    97             (Tactic.rtac (@{thm typedef_cpo} OF cpo_thms) 1);
       
    98         val cpo_thms' = map (Thm.transfer theory') cpo_thms;
       
    99       in
       
   100         theory'
       
   101         |> Sign.add_path (Binding.name_of name)
       
   102         |> PureThy.add_thms
       
   103           ([((Binding.prefix_name "adm_" name, admissible'), []),
       
   104             ((Binding.prefix_name "cont_" Rep_name, @{thm typedef_cont_Rep} OF cpo_thms'), []),
       
   105             ((Binding.prefix_name "cont_" Abs_name, @{thm typedef_cont_Abs} OF cpo_thms'), []),
       
   106             ((Binding.prefix_name "lub_" name, @{thm typedef_lub} OF cpo_thms'), []),
       
   107             ((Binding.prefix_name "thelub_" name, @{thm typedef_thelub} OF cpo_thms'), []),
       
   108             ((Binding.prefix_name "compact_" name, @{thm typedef_compact} OF cpo_thms'), [])])
       
   109         |> snd
       
   110         |> Sign.parent_path
       
   111       end;
       
   112 
       
   113     fun make_pcpo UU_mem (type_def, below_def, set_def) theory =
       
   114       let
       
   115         val UU_mem' = fold_rule (the_list set_def) UU_mem;
       
   116         val pcpo_thms = map (Thm.transfer theory) [type_def, below_def, UU_mem'];
       
   117         val theory' = theory
       
   118           |> AxClass.prove_arity (full_tname, lhs_sorts, @{sort pcpo})
       
   119             (Tactic.rtac (@{thm typedef_pcpo} OF pcpo_thms) 1);
       
   120         val pcpo_thms' = map (Thm.transfer theory') pcpo_thms;
       
   121       in
       
   122         theory'
       
   123         |> Sign.add_path (Binding.name_of name)
       
   124         |> PureThy.add_thms
       
   125           ([((Binding.suffix_name "_strict" Rep_name, @{thm typedef_Rep_strict} OF pcpo_thms'), []),
       
   126             ((Binding.suffix_name "_strict" Abs_name, @{thm typedef_Abs_strict} OF pcpo_thms'), []),
       
   127             ((Binding.suffix_name "_strict_iff" Rep_name, @{thm typedef_Rep_strict_iff} OF pcpo_thms'), []),
       
   128             ((Binding.suffix_name "_strict_iff" Abs_name, @{thm typedef_Abs_strict_iff} OF pcpo_thms'), []),
       
   129             ((Binding.suffix_name "_defined" Rep_name, @{thm typedef_Rep_defined} OF pcpo_thms'), []),
       
   130             ((Binding.suffix_name "_defined" Abs_name, @{thm typedef_Abs_defined} OF pcpo_thms'), [])])
       
   131         |> snd
       
   132         |> Sign.parent_path
       
   133       end;
       
   134 
       
   135     fun pcpodef_result UU_mem admissible =
       
   136       make_po (Tactic.rtac exI 1 THEN Tactic.rtac UU_mem 1)
       
   137       #-> (fn defs => make_cpo admissible defs #> make_pcpo UU_mem defs);
       
   138 
       
   139     fun cpodef_result nonempty admissible =
       
   140       make_po (Tactic.rtac nonempty 1)
       
   141       #-> make_cpo admissible;
       
   142   in
       
   143     if pcpo
       
   144     then (goal_UU_mem, goal_admissible, pcpodef_result)
       
   145     else (goal_nonempty, goal_admissible, cpodef_result)
       
   146   end
       
   147   handle ERROR msg =>
       
   148     cat_error msg ("The error(s) above occurred in cpodef " ^ quote (Binding.str_of name));
       
   149 
       
   150 
       
   151 (* proof interface *)
       
   152 
       
   153 local
       
   154 
       
   155 fun gen_pcpodef_proof prep_term pcpo ((def, name), typ, set, opt_morphs) thy =
       
   156   let
       
   157     val (goal1, goal2, make_result) =
       
   158       prepare_pcpodef prep_term pcpo def name typ set opt_morphs thy;
       
   159     fun after_qed [[th1, th2]] = ProofContext.theory (make_result th1 th2);
       
   160   in Proof.theorem_i NONE after_qed [[(goal1, []), (goal2, [])]] (ProofContext.init thy) end;
       
   161 
       
   162 in
       
   163 
       
   164 fun pcpodef_proof x = gen_pcpodef_proof Syntax.check_term true x;
       
   165 fun pcpodef_proof_cmd x = gen_pcpodef_proof Syntax.read_term true x;
       
   166 
       
   167 fun cpodef_proof x = gen_pcpodef_proof Syntax.check_term false x;
       
   168 fun cpodef_proof_cmd x = gen_pcpodef_proof Syntax.read_term false x;
       
   169 
       
   170 end;
       
   171 
       
   172 
       
   173 
       
   174 (** outer syntax **)
       
   175 
       
   176 local structure P = OuterParse and K = OuterKeyword in
       
   177 
       
   178 val typedef_proof_decl =
       
   179   Scan.optional (P.$$$ "(" |--
       
   180       ((P.$$$ "open" >> K false) -- Scan.option P.binding || P.binding >> (fn s => (true, SOME s)))
       
   181         --| P.$$$ ")") (true, NONE) --
       
   182     (P.type_args -- P.binding) -- P.opt_infix -- (P.$$$ "=" |-- P.term) --
       
   183     Scan.option (P.$$$ "morphisms" |-- P.!!! (P.binding -- P.binding));
       
   184 
       
   185 fun mk_pcpodef_proof pcpo ((((((def, opt_name), (vs, t)), mx), A), morphs)) =
       
   186   (if pcpo then pcpodef_proof_cmd else cpodef_proof_cmd)
       
   187     ((def, the_default (Binding.map_name (Syntax.type_name mx) t) opt_name), (t, vs, mx), A, morphs);
       
   188 
       
   189 val _ =
       
   190   OuterSyntax.command "pcpodef" "HOLCF type definition (requires admissibility proof)" K.thy_goal
       
   191     (typedef_proof_decl >>
       
   192       (Toplevel.print oo (Toplevel.theory_to_proof o mk_pcpodef_proof true)));
       
   193 
       
   194 val _ =
       
   195   OuterSyntax.command "cpodef" "HOLCF type definition (requires admissibility proof)" K.thy_goal
       
   196     (typedef_proof_decl >>
       
   197       (Toplevel.print oo (Toplevel.theory_to_proof o mk_pcpodef_proof false)));
       
   198 
       
   199 end;
       
   200 
       
   201 end;