src/HOLCF/Tools/Domain/domain_axioms.ML
author huffman
Thu, 19 Nov 2009 21:06:22 -0800
changeset 33807 ce8d2e8bca21
parent 33802 48ce3a1063f2
child 33809 033831fd9ef3
permissions -rw-r--r--
domain_isomorphism package defines combined copy function
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 ->
31288
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
    14
      string -> Domain_Library.eq list -> int -> Domain_Library.eq ->
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 ->
31288
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
    19
      bstring -> 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
    20
end;
dcf8a7a66bd1 make domain package ML interface more consistent with datatype package; use binding instead of bstring
huffman
parents: 30912
diff changeset
    21
dcf8a7a66bd1 make domain package ML interface more consistent with datatype package; use binding instead of bstring
huffman
parents: 30912
diff changeset
    22
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 32952
diff changeset
    23
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
    24
struct
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    25
31227
0aa6afd229d3 indentation; export Domain_Axioms.calc_axioms
huffman
parents: 31023
diff changeset
    26
open Domain_Library;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    27
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    28
infixr 0 ===>;infixr 0 ==>;infix 0 == ; 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    29
infix 1 ===; infix 1 ~= ; infix 1 <<; infix 1 ~<<;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    30
infix 9 `   ; infix 9 `% ; infix 9 `%%; infixr 9 oo;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    31
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 31227
diff changeset
    32
(* 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
    33
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
    34
    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
    35
                 (@{type_name "++"}, @{const_name "ssum_map"}),
b4210cc3ac97 map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents: 33396
diff changeset
    36
                 (@{type_name "**"}, @{const_name "sprod_map"}),
b4210cc3ac97 map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents: 33396
diff changeset
    37
                 (@{type_name "*"}, @{const_name "cprod_map"}),
b4210cc3ac97 map functions for various types, with ep_pair/deflation/finite_deflation lemmas
huffman
parents: 33396
diff changeset
    38
                 (@{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
    39
33801
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
    40
fun copy_of_dtyp tab r dt =
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
    41
    if DatatypeAux.is_rec_type dt then copy tab r dt else ID
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
    42
and copy tab r (DatatypeAux.DtRec i) = r i
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
    43
  | copy tab r (DatatypeAux.DtTFree a) = ID
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
    44
  | copy tab r (DatatypeAux.DtType (c, ds)) =
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
    45
    case Symtab.lookup tab c of
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
    46
      SOME f => list_ccomb (%%:f, map (copy_of_dtyp tab r) ds)
31288
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
    47
    | 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
    48
31227
0aa6afd229d3 indentation; export Domain_Axioms.calc_axioms
huffman
parents: 31023
diff changeset
    49
fun calc_axioms
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    50
    (definitional : bool)
33801
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
    51
    (map_tab : string Symtab.table)
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    52
    (comp_dname : string)
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    53
    (eqs : eq list)
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    54
    (n : int)
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    55
    (eqn as ((dname,_),cons) : eq)
31227
0aa6afd229d3 indentation; export Domain_Axioms.calc_axioms
huffman
parents: 31023
diff changeset
    56
    : string * (string * term) list * (string * term) list =
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    57
  let
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    58
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    59
(* ----- axioms and definitions concerning the isomorphism ------------------ *)
31227
0aa6afd229d3 indentation; export Domain_Axioms.calc_axioms
huffman
parents: 31023
diff changeset
    60
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    61
    val dc_abs = %%:(dname^"_abs");
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    62
    val dc_rep = %%:(dname^"_rep");
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    63
    val x_name'= "x";
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    64
    val x_name = idx_name eqs x_name' (n+1);
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    65
    val dnam = Long_Name.base_name dname;
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    66
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    67
    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
    68
    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
    69
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    70
    val when_def = ("when_def",%%:(dname^"_when") == 
33801
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
    71
        List.foldr (uncurry /\ ) (/\x_name'((when_body cons (fn (x,y) =>
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
    72
          Bound(1+length cons+x-y)))`(dc_rep`Bound 0))) (when_funs cons));
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
    73
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    74
    val copy_def =
33801
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
    75
      let fun r i = proj (Bound 0) eqs i;
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
    76
      in
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
    77
        ("copy_def", %%:(dname^"_copy") == /\ "f"
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
    78
          (dc_abs oo (copy_of_dtyp map_tab r (dtyp_of_eq eqn)) oo dc_rep))
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
    79
      end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    80
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    81
(* -- definitions concerning the constructors, discriminators and selectors - *)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    82
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    83
    fun con_def m n (_,args) = let
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    84
      fun idxs z x arg = (if is_lazy arg then mk_up else I) (Bound(z-x));
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    85
      fun parms vs = mk_stuple (mapn (idxs(length vs)) 1 vs);
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    86
      fun inj y 1 _ = y
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    87
        | inj y _ 0 = mk_sinl y
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    88
        | inj y i j = mk_sinr (inj y (i-1) (j-1));
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    89
    in List.foldr /\# (dc_abs`(inj (parms args) m n)) args end;
31288
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
    90
          
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    91
    val con_defs = mapn (fn n => fn (con,args) =>
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    92
                                    (extern_name con ^"_def", %%:con == con_def (length cons) n (con,args))) 0 cons;
31288
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
    93
          
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    94
    val dis_defs = let
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    95
      fun ddef (con,_) = (dis_name con ^"_def",%%:(dis_name con) == 
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    96
                                              list_ccomb(%%:(dname^"_when"),map 
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    97
                                                                              (fn (con',args) => (List.foldr /\#
31288
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
    98
      (if con'=con then TT else FF) args)) cons))
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
    99
    in map ddef cons end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   100
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   101
    val mat_defs =
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
        fun mdef (con,_) =
31288
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
   104
          let
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   105
            val k = Bound 0
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   106
            val x = Bound 1
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   107
            fun one_con (con', args') =
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   108
                if con'=con then k else List.foldr /\# mk_fail args'
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   109
            val w = list_ccomb(%%:(dname^"_when"), map one_con cons)
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   110
            val rhs = /\ "x" (/\ "k" (w ` x))
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   111
          in (mat_name con ^"_def", %%:(mat_name con) == rhs) end
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   112
      in map mdef cons end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   113
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   114
    val pat_defs =
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   115
      let
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   116
        fun pdef (con,args) =
31288
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
   117
          let
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   118
            val ps = mapn (fn n => fn _ => %:("pat" ^ string_of_int n)) 1 args;
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   119
            val xs = map (bound_arg args) args;
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   120
            val r = Bound (length args);
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   121
            val rhs = case args of [] => mk_return HOLogic.unit
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   122
                                 | _ => mk_ctuple_pat ps ` mk_ctuple xs;
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   123
            fun one_con (con',args') = List.foldr /\# (if con'=con then rhs else mk_fail) args';
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   124
          in (pat_name con ^"_def", list_comb (%%:(pat_name con), ps) == 
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   125
                                              list_ccomb(%%:(dname^"_when"), map one_con cons))
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   126
          end
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   127
      in map pdef cons end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   128
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   129
    val sel_defs = let
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   130
      fun sdef con n arg = Option.map (fn sel => (sel^"_def",%%:sel == 
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   131
                                                            list_ccomb(%%:(dname^"_when"),map 
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   132
                                                                                            (fn (con',args) => if con'<>con then UU else
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   133
                                                                                                               List.foldr /\# (Bound (length args - n)) args) cons))) (sel_of arg);
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   134
    in map_filter I (maps (fn (con,args) => mapn (sdef con) 1 args) cons) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   135
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   136
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   137
(* ----- axiom and definitions concerning induction ------------------------- *)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   138
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   139
    val reach_ax = ("reach", mk_trp(proj (mk_fix (%%:(comp_dname^"_copy"))) eqs n
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   140
                                         `%x_name === %:x_name));
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   141
    val take_def =
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   142
        ("take_def",
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   143
         %%:(dname^"_take") ==
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   144
            mk_lam("n",proj
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   145
                         (mk_iterate (Bound 0, %%:(comp_dname^"_copy"), UU)) eqs n));
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   146
    val finite_def =
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   147
        ("finite_def",
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   148
         %%:(dname^"_finite") ==
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   149
            mk_lam(x_name,
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   150
                   mk_ex("n",(%%:(dname^"_take") $ Bound 0)`Bound 1 === Bound 1)));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   151
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   152
  in (dnam,
33802
48ce3a1063f2 domain_isomorphism package defines copy functions
huffman
parents: 33801
diff changeset
   153
      (if definitional then [reach_ax] else [abs_iso_ax, rep_iso_ax, reach_ax]),
48ce3a1063f2 domain_isomorphism package defines copy functions
huffman
parents: 33801
diff changeset
   154
      (if definitional then [when_def] else [when_def, copy_def]) @
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   155
      con_defs @ dis_defs @ mat_defs @ pat_defs @ sel_defs @
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   156
      [take_def, finite_def])
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   157
  end; (* let (calc_axioms) *)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   158
30483
0c398040969c added legacy type inference (from fixrec_package.ML);
wenzelm
parents: 30364
diff changeset
   159
0c398040969c added legacy type inference (from fixrec_package.ML);
wenzelm
parents: 30364
diff changeset
   160
(* legacy type inference *)
0c398040969c added legacy type inference (from fixrec_package.ML);
wenzelm
parents: 30364
diff changeset
   161
0c398040969c added legacy type inference (from fixrec_package.ML);
wenzelm
parents: 30364
diff changeset
   162
fun legacy_infer_term thy t =
31227
0aa6afd229d3 indentation; export Domain_Axioms.calc_axioms
huffman
parents: 31023
diff changeset
   163
    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
   164
0c398040969c added legacy type inference (from fixrec_package.ML);
wenzelm
parents: 30364
diff changeset
   165
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
   166
0c398040969c added legacy type inference (from fixrec_package.ML);
wenzelm
parents: 30364
diff changeset
   167
fun infer_props thy = map (apsnd (legacy_infer_prop thy));
0c398040969c added legacy type inference (from fixrec_package.ML);
wenzelm
parents: 30364
diff changeset
   168
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   169
29585
c23295521af5 binding replaces bstring
haftmann
parents: 28965
diff changeset
   170
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
   171
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
   172
29585
c23295521af5 binding replaces bstring
haftmann
parents: 28965
diff changeset
   173
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
   174
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
   175
30131
6be1be402ef0 use TheoryData to keep track of pattern match combinators
huffman
parents: 29585
diff changeset
   176
fun add_matchers (((dname,_),cons) : eq) thy =
31227
0aa6afd229d3 indentation; export Domain_Axioms.calc_axioms
huffman
parents: 31023
diff changeset
   177
    let
31288
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
   178
      val con_names = map fst cons;
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
   179
      val mat_names = map mat_name con_names;
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
   180
      fun qualify n = Sign.full_name thy (Binding.name n);
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
   181
      val ms = map qualify con_names ~~ map qualify mat_names;
31738
7b9b9ba532ca discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31288
diff changeset
   182
    in Fixrec.add_matchers ms thy end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   183
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   184
fun add_axioms definitional comp_dnam (eqs : eq list) thy' =
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   185
  let
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   186
    val comp_dname = Sign.full_bname thy' comp_dnam;
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   187
    val dnames = map (fst o fst) eqs;
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   188
    val x_name = idx_name dnames "x"; 
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   189
    fun copy_app dname = %%:(dname^"_copy")`Bound 0;
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   190
    val copy_def = ("copy_def" , %%:(comp_dname^"_copy") ==
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   191
                                 /\ "f"(mk_ctuple (map copy_app dnames)));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   192
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   193
    fun one_con (con,args) =
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   194
      let
31288
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
   195
        val nonrec_args = filter_out is_rec args;
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 33245
diff changeset
   196
        val    rec_args = filter is_rec args;
31288
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
   197
        val    recs_cnt = length rec_args;
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
   198
        val allargs     = nonrec_args @ rec_args
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
   199
                          @ map (upd_vname (fn s=> s^"'")) rec_args;
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
   200
        val allvns      = map vname allargs;
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
   201
        fun vname_arg s arg = if is_rec arg then vname arg^s else vname arg;
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
   202
        val vns1        = map (vname_arg "" ) args;
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
   203
        val vns2        = map (vname_arg "'") args;
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
   204
        val allargs_cnt = length nonrec_args + 2*recs_cnt;
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
   205
        val rec_idxs    = (recs_cnt-1) downto 0;
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
   206
        val nonlazy_idxs = map snd (filter_out (fn (arg,_) => is_lazy arg)
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
   207
                                               (allargs~~((allargs_cnt-1) downto 0)));
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
   208
        fun rel_app i ra = proj (Bound(allargs_cnt+2)) eqs (rec_of ra) $ 
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
   209
                                Bound (2*recs_cnt-i) $ Bound (recs_cnt-i);
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
   210
        val capps =
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   211
          List.foldr
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   212
            mk_conj
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   213
            (mk_conj(
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   214
             Bound(allargs_cnt+1)===list_ccomb(%%:con,map (bound_arg allvns) vns1),
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   215
             Bound(allargs_cnt+0)===list_ccomb(%%:con,map (bound_arg allvns) vns2)))
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   216
            (mapn rel_app 1 rec_args);
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   217
      in
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   218
        List.foldr
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   219
          mk_ex
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   220
          (Library.foldr mk_conj
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   221
                         (map (defined o Bound) nonlazy_idxs,capps)) allvns
31288
67dff9c5b2bd remove hard tabs, fix indentation
huffman
parents: 31232
diff changeset
   222
      end;
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   223
    fun one_comp n (_,cons) =
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   224
        mk_all (x_name(n+1),
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   225
        mk_all (x_name(n+1)^"'",
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   226
        mk_imp (proj (Bound 2) eqs n $ Bound 1 $ Bound 0,
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   227
        foldr1 mk_disj (mk_conj(Bound 1 === UU,Bound 0 === UU)
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   228
                        ::map one_con cons))));
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   229
    val bisim_def =
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   230
        ("bisim_def", %%:(comp_dname^"_bisim") ==
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   231
                         mk_lam("R", foldr1 mk_conj (mapn one_comp 0 eqs)));
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   232
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   233
    fun add_one (dnam, axs, dfs) =
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   234
        Sign.add_path dnam
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 32952
diff changeset
   235
          #> add_defs_infer dfs
65232054ffd0 eliminated some old folds;
wenzelm
parents: 32952
diff changeset
   236
          #> add_axioms_infer axs
65232054ffd0 eliminated some old folds;
wenzelm
parents: 32952
diff changeset
   237
          #> Sign.parent_path;
31227
0aa6afd229d3 indentation; export Domain_Axioms.calc_axioms
huffman
parents: 31023
diff changeset
   238
33801
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
   239
    val map_tab = Domain_Isomorphism.get_map_tab thy';
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
   240
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   241
    val thy = thy'
33801
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33798
diff changeset
   242
      |> fold add_one (mapn (calc_axioms definitional map_tab comp_dname eqs) 0 eqs);
31227
0aa6afd229d3 indentation; export Domain_Axioms.calc_axioms
huffman
parents: 31023
diff changeset
   243
33807
ce8d2e8bca21 domain_isomorphism package defines combined copy function
huffman
parents: 33802
diff changeset
   244
    val use_copy_def = length eqs>1 andalso not definitional;
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   245
  in
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   246
    thy
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   247
    |> Sign.add_path comp_dnam  
33807
ce8d2e8bca21 domain_isomorphism package defines combined copy function
huffman
parents: 33802
diff changeset
   248
    |> add_defs_infer (bisim_def::(if use_copy_def then [copy_def] else []))
33798
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   249
    |> Sign.parent_path
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   250
    |> fold add_matchers eqs
46cbbcbd4e68 clean up indentation; add 'definitional' option flag
huffman
parents: 33504
diff changeset
   251
  end; (* let (add_axioms) *)
31227
0aa6afd229d3 indentation; export Domain_Axioms.calc_axioms
huffman
parents: 31023
diff changeset
   252
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   253
end; (* struct *)