src/HOLCF/Tools/Domain/domain_axioms.ML
author huffman
Tue, 02 Mar 2010 04:31:50 -0800
changeset 35497 979706bd5c16
parent 35495 dc59e781669f
child 35499 6acef0aea07d
permissions -rw-r--r--
re-enable bisim code, now in domain_theorems.ML
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32126
a5042f260440 obey captialized directory names convention
haftmann
parents: 31738
diff changeset
     1
(*  Title:      HOLCF/Tools/Domain/domain_axioms.ML
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     2
    Author:     David von Oheimb
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     3
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     4
Syntax generator for domain command.
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     5
*)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     6
30919
dcf8a7a66bd1 make domain package ML interface more consistent with datatype package; use binding instead of bstring
huffman
parents: 30912
diff changeset
     7
signature DOMAIN_AXIOMS =
dcf8a7a66bd1 make domain package ML interface more consistent with datatype package; use binding instead of bstring
huffman
parents: 30912
diff changeset
     8
sig
33801
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
     9
  val copy_of_dtyp :
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
    10
      string Symtab.table -> (int -> term) -> Datatype.dtyp -> term
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31227
diff changeset
    11
31288
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
    12
  val calc_axioms :
33801
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
    13
      bool -> string Symtab.table ->
35497
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35495
diff changeset
    14
      Domain_Library.eq list -> int -> Domain_Library.eq ->
31288
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
    15
      string * (string * term) list * (string * term) list
31227
0aa6afd229d3 indentation; export Domain_Axioms.calc_axioms
huffman
parents: 31023
diff changeset
    16
31288
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
    17
  val add_axioms :
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    18
      bool ->
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
    19
      ((string * typ list) *
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
    20
       (binding * (bool * binding option * typ) list * mixfix) list) list ->
35497
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35495
diff changeset
    21
      Domain_Library.eq list -> theory -> theory
30919
dcf8a7a66bd1 make domain package ML interface more consistent with datatype package; use binding instead of bstring
huffman
parents: 30912
diff changeset
    22
end;
dcf8a7a66bd1 make domain package ML interface more consistent with datatype package; use binding instead of bstring
huffman
parents: 30912
diff changeset
    23
dcf8a7a66bd1 make domain package ML interface more consistent with datatype package; use binding instead of bstring
huffman
parents: 30912
diff changeset
    24
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 32952
diff changeset
    25
structure Domain_Axioms : DOMAIN_AXIOMS =
30919
dcf8a7a66bd1 make domain package ML interface more consistent with datatype package; use binding instead of bstring
huffman
parents: 30912
diff changeset
    26
struct
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    27
31227
0aa6afd229d3 indentation; export Domain_Axioms.calc_axioms
huffman
parents: 31023
diff changeset
    28
open Domain_Library;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    29
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    30
infixr 0 ===>;infixr 0 ==>;infix 0 == ; 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    31
infix 1 ===; infix 1 ~= ; infix 1 <<; infix 1 ~<<;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    32
infix 9 `   ; infix 9 `% ; infix 9 `%%; infixr 9 oo;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    33
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31227
diff changeset
    34
(* FIXME: use theory data for this *)
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31227
diff changeset
    35
val copy_tab : string Symtab.table =
33504
b4210cc3ac97 map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents: 33396
diff changeset
    36
    Symtab.make [(@{type_name "->"}, @{const_name "cfun_map"}),
b4210cc3ac97 map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents: 33396
diff changeset
    37
                 (@{type_name "++"}, @{const_name "ssum_map"}),
b4210cc3ac97 map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents: 33396
diff changeset
    38
                 (@{type_name "**"}, @{const_name "sprod_map"}),
b4210cc3ac97 map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents: 33396
diff changeset
    39
                 (@{type_name "*"}, @{const_name "cprod_map"}),
b4210cc3ac97 map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents: 33396
diff changeset
    40
                 (@{type_name "u"}, @{const_name "u_map"})];
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31227
diff changeset
    41
33801
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
    42
fun copy_of_dtyp tab r dt =
33971
9c7fa7f76950 modernized structure Datatype_Aux
haftmann
parents: 33809
diff changeset
    43
    if Datatype_Aux.is_rec_type dt then copy tab r dt else ID
9c7fa7f76950 modernized structure Datatype_Aux
haftmann
parents: 33809
diff changeset
    44
and copy tab r (Datatype_Aux.DtRec i) = r i
9c7fa7f76950 modernized structure Datatype_Aux
haftmann
parents: 33809
diff changeset
    45
  | copy tab r (Datatype_Aux.DtTFree a) = ID
9c7fa7f76950 modernized structure Datatype_Aux
haftmann
parents: 33809
diff changeset
    46
  | copy tab r (Datatype_Aux.DtType (c, ds)) =
33801
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
    47
    case Symtab.lookup tab c of
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
    48
      SOME f => list_ccomb (%%:f, map (copy_of_dtyp tab r) ds)
31288
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
    49
    | NONE => (warning ("copy_of_dtyp: unknown type constructor " ^ c); ID);
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31227
diff changeset
    50
31227
0aa6afd229d3 indentation; export Domain_Axioms.calc_axioms
huffman
parents: 31023
diff changeset
    51
fun calc_axioms
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    52
    (definitional : bool)
33801
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
    53
    (map_tab : string Symtab.table)
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    54
    (eqs : eq list)
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    55
    (n : int)
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    56
    (eqn as ((dname,_),cons) : eq)
31227
0aa6afd229d3 indentation; export Domain_Axioms.calc_axioms
huffman
parents: 31023
diff changeset
    57
    : string * (string * term) list * (string * term) list =
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    58
  let
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    59
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    60
(* ----- axioms and definitions concerning the isomorphism ------------------ *)
31227
0aa6afd229d3 indentation; export Domain_Axioms.calc_axioms
huffman
parents: 31023
diff changeset
    61
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    62
    val dc_abs = %%:(dname^"_abs");
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    63
    val dc_rep = %%:(dname^"_rep");
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    64
    val x_name'= "x";
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    65
    val x_name = idx_name eqs x_name' (n+1);
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    66
    val dnam = Long_Name.base_name dname;
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    67
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    68
    val abs_iso_ax = ("abs_iso", mk_trp(dc_rep`(dc_abs`%x_name') === %:x_name'));
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    69
    val rep_iso_ax = ("rep_iso", mk_trp(dc_abs`(dc_rep`%x_name') === %:x_name'));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    70
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    71
(* ----- axiom and definitions concerning induction ------------------------- *)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    72
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    73
    val finite_def =
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    74
        ("finite_def",
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    75
         %%:(dname^"_finite") ==
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    76
            mk_lam(x_name,
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    77
                   mk_ex("n",(%%:(dname^"_take") $ Bound 0)`Bound 1 === Bound 1)));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    78
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    79
  in (dnam,
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
    80
      (if definitional then [] else [abs_iso_ax, rep_iso_ax]),
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
    81
      [finite_def])
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    82
  end; (* let (calc_axioms) *)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    83
30483
0c398040969c added legacy type inference (from fixrec_package.ML);
wenzelm
parents: 30364
diff changeset
    84
0c398040969c added legacy type inference (from fixrec_package.ML);
wenzelm
parents: 30364
diff changeset
    85
(* legacy type inference *)
0c398040969c added legacy type inference (from fixrec_package.ML);
wenzelm
parents: 30364
diff changeset
    86
0c398040969c added legacy type inference (from fixrec_package.ML);
wenzelm
parents: 30364
diff changeset
    87
fun legacy_infer_term thy t =
31227
0aa6afd229d3 indentation; export Domain_Axioms.calc_axioms
huffman
parents: 31023
diff changeset
    88
    singleton (Syntax.check_terms (ProofContext.init thy)) (Sign.intern_term thy t);
30483
0c398040969c added legacy type inference (from fixrec_package.ML);
wenzelm
parents: 30364
diff changeset
    89
0c398040969c added legacy type inference (from fixrec_package.ML);
wenzelm
parents: 30364
diff changeset
    90
fun legacy_infer_prop thy t = legacy_infer_term thy (TypeInfer.constrain propT t);
0c398040969c added legacy type inference (from fixrec_package.ML);
wenzelm
parents: 30364
diff changeset
    91
0c398040969c added legacy type inference (from fixrec_package.ML);
wenzelm
parents: 30364
diff changeset
    92
fun infer_props thy = map (apsnd (legacy_infer_prop thy));
0c398040969c added legacy type inference (from fixrec_package.ML);
wenzelm
parents: 30364
diff changeset
    93
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    94
29585
c23295521af5 binding replaces bstring
haftmann
parents: 28965
diff changeset
    95
fun add_axioms_i x = snd o PureThy.add_axioms (map (Thm.no_attributes o apfst Binding.name) x);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    96
fun add_axioms_infer axms thy = add_axioms_i (infer_props thy axms) thy;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    97
29585
c23295521af5 binding replaces bstring
haftmann
parents: 28965
diff changeset
    98
fun add_defs_i x = snd o (PureThy.add_defs false) (map (Thm.no_attributes o apfst Binding.name) x);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    99
fun add_defs_infer defs thy = add_defs_i (infer_props thy defs) thy;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   100
35497
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35495
diff changeset
   101
fun add_axioms definitional eqs' (eqs : eq list) thy' =
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   102
  let
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   103
    val dnames = map (fst o fst) eqs;
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   104
    val x_name = idx_name dnames "x"; 
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   105
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   106
    fun add_one (dnam, axs, dfs) =
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   107
        Sign.add_path dnam
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 32952
diff changeset
   108
          #> add_axioms_infer axs
65232054ffd0 eliminated some old folds;
wenzelm
parents: 32952
diff changeset
   109
          #> Sign.parent_path;
31227
0aa6afd229d3 indentation; export Domain_Axioms.calc_axioms
huffman
parents: 31023
diff changeset
   110
33801
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
   111
    val map_tab = Domain_Isomorphism.get_map_tab thy';
35497
979706bd5c16 re-enable bisim code, now in domain_theorems.ML
huffman
parents: 35495
diff changeset
   112
    val axs = mapn (calc_axioms definitional map_tab eqs) 0 eqs;
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   113
    val thy = thy' |> fold add_one axs;
33801
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
   114
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   115
    fun get_iso_info ((dname, tyvars), cons') =
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   116
      let
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   117
        fun opt_lazy (lazy,_,t) = if lazy then mk_uT t else t
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   118
        fun prod     (_,args,_) =
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   119
            case args of [] => oneT
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   120
                       | _ => foldr1 mk_sprodT (map opt_lazy args);
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   121
        val ax_abs_iso = PureThy.get_thm thy (dname ^ ".abs_iso");
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   122
        val ax_rep_iso = PureThy.get_thm thy (dname ^ ".rep_iso");
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   123
        val lhsT = Type(dname,tyvars);
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   124
        val rhsT = foldr1 mk_ssumT (map prod cons');
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   125
        val rep_const = Const(dname^"_rep", lhsT ->> rhsT);
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   126
        val abs_const = Const(dname^"_abs", rhsT ->> lhsT);
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   127
      in
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   128
        {
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   129
          absT = lhsT,
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   130
          repT = rhsT,
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   131
          abs_const = abs_const,
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   132
          rep_const = rep_const,
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   133
          abs_inverse = ax_abs_iso,
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   134
          rep_inverse = ax_rep_iso
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   135
        }
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   136
      end;
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   137
    val dom_binds = map (Binding.name o Long_Name.base_name) dnames;
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   138
    val thy =
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   139
        if definitional then thy
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   140
        else snd (Domain_Isomorphism.define_take_functions
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   141
                    (dom_binds ~~ map get_iso_info eqs') thy);
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   142
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   143
    fun add_one' (dnam, axs, dfs) =
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   144
        Sign.add_path dnam
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   145
          #> add_defs_infer dfs
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   146
          #> Sign.parent_path;
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   147
    val thy = fold add_one' axs thy;
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   148
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   149
    (* declare lub_take axioms *)
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   150
    local
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   151
      fun ax_lub_take dname =
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   152
        let
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   153
          val dnam : string = Long_Name.base_name dname;
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   154
          val take_const = %%:(dname^"_take");
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   155
          val lub = %%: @{const_name lub};
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   156
          val image = %%: @{const_name image};
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   157
          val UNIV = %%: @{const_name UNIV};
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   158
          val lhs = lub $ (image $ take_const $ UNIV);
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   159
          val ax = mk_trp (lhs === ID);
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   160
        in
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   161
          add_one (dnam, [("lub_take", ax)], [])
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   162
        end
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   163
    in
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   164
      val thy =
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   165
          if definitional then thy
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   166
          else fold ax_lub_take dnames thy
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35486
diff changeset
   167
    end;
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   168
  in
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   169
    thy
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   170
  end; (* let (add_axioms) *)
31227
0aa6afd229d3 indentation; export Domain_Axioms.calc_axioms
huffman
parents: 31023
diff changeset
   171
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   172
end; (* struct *)