src/HOLCF/Tools/domain/domain_axioms.ML
author blanchet
Wed, 04 Mar 2009 10:45:52 +0100
changeset 30240 5b25fee0362c
parent 29585 c23295521af5
child 30242 aea5d7fa7ef5
permissions -rw-r--r--
Merge.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOLCF/Tools/domain/domain_axioms.ML
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
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     7
structure Domain_Axioms = struct
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     8
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     9
local
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    10
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    11
open Domain_Library;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    12
infixr 0 ===>;infixr 0 ==>;infix 0 == ; 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    13
infix 1 ===; infix 1 ~= ; infix 1 <<; infix 1 ~<<;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    14
infix 9 `   ; infix 9 `% ; infix 9 `%%; infixr 9 oo;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    15
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    16
fun calc_axioms comp_dname (eqs : eq list) n (((dname,_),cons) : eq)=
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    17
let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    18
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    19
(* ----- axioms and definitions concerning the isomorphism ------------------ *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    20
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    21
  val dc_abs = %%:(dname^"_abs");
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    22
  val dc_rep = %%:(dname^"_rep");
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    23
  val x_name'= "x";
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    24
  val x_name = idx_name eqs x_name' (n+1);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    25
  val dnam = Sign.base_name dname;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    26
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    27
  val abs_iso_ax = ("abs_iso", mk_trp(dc_rep`(dc_abs`%x_name') === %:x_name'));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    28
  val rep_iso_ax = ("rep_iso", mk_trp(dc_abs`(dc_rep`%x_name') === %:x_name'));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    29
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    30
  val when_def = ("when_def",%%:(dname^"_when") == 
30240
blanchet
parents: 29585
diff changeset
    31
     List.foldr (uncurry /\ ) (/\x_name'((when_body cons (fn (x,y) =>
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    32
				Bound(1+length cons+x-y)))`(dc_rep`Bound 0))) (when_funs cons));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    33
  
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    34
  val copy_def = let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    35
    fun idxs z x arg = if is_rec arg
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    36
			 then (cproj (Bound z) eqs (rec_of arg))`Bound(z-x)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    37
			 else Bound(z-x);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    38
    fun one_con (con,args) =
30240
blanchet
parents: 29585
diff changeset
    39
        List.foldr /\# (list_ccomb (%%:con, mapn (idxs (length args)) 1 args)) args;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    40
  in ("copy_def", %%:(dname^"_copy") ==
30240
blanchet
parents: 29585
diff changeset
    41
       /\ "f" (list_ccomb (%%:(dname^"_when"), map one_con cons))) end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    42
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    43
(* -- definitions concerning the constructors, discriminators and selectors - *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    44
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    45
  fun con_def m n (_,args) = let
26012
f6917792f8a4 new term-building combinators
huffman
parents: 24712
diff changeset
    46
    fun idxs z x arg = (if is_lazy arg then mk_up else I) (Bound(z-x));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    47
    fun parms vs = mk_stuple (mapn (idxs(length vs)) 1 vs);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    48
    fun inj y 1 _ = y
26012
f6917792f8a4 new term-building combinators
huffman
parents: 24712
diff changeset
    49
    |   inj y _ 0 = mk_sinl y
f6917792f8a4 new term-building combinators
huffman
parents: 24712
diff changeset
    50
    |   inj y i j = mk_sinr (inj y (i-1) (j-1));
30240
blanchet
parents: 29585
diff changeset
    51
  in List.foldr /\# (dc_abs`(inj (parms args) m n)) args end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    52
  
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    53
  val con_defs = mapn (fn n => fn (con,args) =>
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    54
    (extern_name con ^"_def", %%:con == con_def (length cons) n (con,args))) 0 cons;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    55
  
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    56
  val dis_defs = let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    57
	fun ddef (con,_) = (dis_name con ^"_def",%%:(dis_name con) == 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    58
		 list_ccomb(%%:(dname^"_when"),map 
30240
blanchet
parents: 29585
diff changeset
    59
			(fn (con',args) => (List.foldr /\#
26012
f6917792f8a4 new term-building combinators
huffman
parents: 24712
diff changeset
    60
			   (if con'=con then TT else FF) args)) cons))
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    61
	in map ddef cons end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    62
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    63
  val mat_defs = let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    64
	fun mdef (con,_) = (mat_name con ^"_def",%%:(mat_name con) == 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    65
		 list_ccomb(%%:(dname^"_when"),map 
30240
blanchet
parents: 29585
diff changeset
    66
			(fn (con',args) => (List.foldr /\#
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    67
			   (if con'=con
26012
f6917792f8a4 new term-building combinators
huffman
parents: 24712
diff changeset
    68
                               then mk_return (mk_ctuple (map (bound_arg args) args))
f6917792f8a4 new term-building combinators
huffman
parents: 24712
diff changeset
    69
                               else mk_fail) args)) cons))
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    70
	in map mdef cons end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    71
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    72
  val pat_defs =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    73
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    74
      fun pdef (con,args) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    75
        let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    76
          val ps = mapn (fn n => fn _ => %:("pat" ^ string_of_int n)) 1 args;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    77
          val xs = map (bound_arg args) args;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    78
          val r = Bound (length args);
26012
f6917792f8a4 new term-building combinators
huffman
parents: 24712
diff changeset
    79
          val rhs = case args of [] => mk_return HOLogic.unit
f6917792f8a4 new term-building combinators
huffman
parents: 24712
diff changeset
    80
                                | _ => mk_ctuple_pat ps ` mk_ctuple xs;
30240
blanchet
parents: 29585
diff changeset
    81
          fun one_con (con',args') = List.foldr /\# (if con'=con then rhs else mk_fail) args';
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    82
        in (pat_name con ^"_def", list_comb (%%:(pat_name con), ps) == 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    83
               list_ccomb(%%:(dname^"_when"), map one_con cons))
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    84
        end
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    85
    in map pdef cons end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    86
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    87
  val sel_defs = let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    88
	fun sdef con n arg = Option.map (fn sel => (sel^"_def",%%:sel == 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    89
		 list_ccomb(%%:(dname^"_when"),map 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    90
			(fn (con',args) => if con'<>con then UU else
30240
blanchet
parents: 29585
diff changeset
    91
			 List.foldr /\# (Bound (length args - n)) args) cons))) (sel_of arg);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    92
	in List.mapPartial I (List.concat(map (fn (con,args) => mapn (sdef con) 1 args) cons)) end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    93
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    94
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    95
(* ----- axiom and definitions concerning induction ------------------------- *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    96
26012
f6917792f8a4 new term-building combinators
huffman
parents: 24712
diff changeset
    97
  val reach_ax = ("reach", mk_trp(cproj (mk_fix (%%:(comp_dname^"_copy"))) eqs n
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    98
					`%x_name === %:x_name));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    99
  val take_def = ("take_def",%%:(dname^"_take") == mk_lam("n",cproj
26012
f6917792f8a4 new term-building combinators
huffman
parents: 24712
diff changeset
   100
	     (mk_iterate (Bound 0, %%:(comp_dname^"_copy"), UU)) eqs n));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   101
  val finite_def = ("finite_def",%%:(dname^"_finite") == mk_lam(x_name,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   102
	mk_ex("n",(%%:(dname^"_take") $ Bound 0)`Bound 1 === Bound 1)));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   103
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   104
in (dnam,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   105
    [abs_iso_ax, rep_iso_ax, reach_ax],
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   106
    [when_def, copy_def] @
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   107
     con_defs @ dis_defs @ mat_defs @ pat_defs @ sel_defs @
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   108
    [take_def, finite_def])
30240
blanchet
parents: 29585
diff changeset
   109
end; (* let (calc_axioms) *)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   110
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   111
fun infer_props thy = map (apsnd (FixrecPackage.legacy_infer_prop thy));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   112
29585
c23295521af5 binding replaces bstring
haftmann
parents: 28965
diff changeset
   113
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
   114
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
   115
29585
c23295521af5 binding replaces bstring
haftmann
parents: 28965
diff changeset
   116
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
   117
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
   118
30240
blanchet
parents: 29585
diff changeset
   119
fun add_matchers (((dname,_),cons) : eq) thy =
blanchet
parents: 29585
diff changeset
   120
  let
blanchet
parents: 29585
diff changeset
   121
    val con_names = map fst cons;
blanchet
parents: 29585
diff changeset
   122
    val mat_names = map mat_name con_names;
blanchet
parents: 29585
diff changeset
   123
    fun qualify n = Sign.full_name thy (Binding.name n);
blanchet
parents: 29585
diff changeset
   124
    val ms = map qualify con_names ~~ map qualify mat_names;
blanchet
parents: 29585
diff changeset
   125
  in FixrecPackage.add_matchers ms thy end;
blanchet
parents: 29585
diff changeset
   126
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   127
in (* local *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   128
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   129
fun add_axioms (comp_dnam, eqs : eq list) thy' = let
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 27691
diff changeset
   130
  val comp_dname = Sign.full_bname thy' comp_dnam;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   131
  val dnames = map (fst o fst) eqs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   132
  val x_name = idx_name dnames "x"; 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   133
  fun copy_app dname = %%:(dname^"_copy")`Bound 0;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   134
  val copy_def = ("copy_def" , %%:(comp_dname^"_copy") ==
30240
blanchet
parents: 29585
diff changeset
   135
				    /\ "f"(mk_ctuple (map copy_app dnames)));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   136
  val bisim_def = ("bisim_def",%%:(comp_dname^"_bisim")==mk_lam("R",
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   137
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   138
      fun one_con (con,args) = let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   139
	val nonrec_args = filter_out is_rec args;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   140
	val    rec_args = List.filter     is_rec args;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   141
	val    recs_cnt = length rec_args;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   142
	val allargs     = nonrec_args @ rec_args
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   143
				      @ map (upd_vname (fn s=> s^"'")) rec_args;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   144
	val allvns      = map vname allargs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   145
	fun vname_arg s arg = if is_rec arg then vname arg^s else vname arg;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   146
	val vns1        = map (vname_arg "" ) args;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   147
	val vns2        = map (vname_arg "'") args;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   148
	val allargs_cnt = length nonrec_args + 2*recs_cnt;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   149
	val rec_idxs    = (recs_cnt-1) downto 0;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   150
	val nonlazy_idxs = map snd (filter_out (fn (arg,_) => is_lazy arg)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   151
					 (allargs~~((allargs_cnt-1) downto 0)));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   152
	fun rel_app i ra = proj (Bound(allargs_cnt+2)) eqs (rec_of ra) $ 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   153
			   Bound (2*recs_cnt-i) $ Bound (recs_cnt-i);
30240
blanchet
parents: 29585
diff changeset
   154
	val capps = List.foldr mk_conj (mk_conj(
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   155
	   Bound(allargs_cnt+1)===list_ccomb(%%:con,map (bound_arg allvns) vns1),
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   156
	   Bound(allargs_cnt+0)===list_ccomb(%%:con,map (bound_arg allvns) vns2)))
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   157
           (mapn rel_app 1 rec_args);
30240
blanchet
parents: 29585
diff changeset
   158
        in List.foldr mk_ex (Library.foldr mk_conj 
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   159
			      (map (defined o Bound) nonlazy_idxs,capps)) allvns end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   160
      fun one_comp n (_,cons) =mk_all(x_name(n+1),mk_all(x_name(n+1)^"'",mk_imp(
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   161
	 		proj (Bound 2) eqs n $ Bound 1 $ Bound 0,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   162
         		foldr1 mk_disj (mk_conj(Bound 1 === UU,Bound 0 === UU)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   163
					::map one_con cons))));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   164
    in foldr1 mk_conj (mapn one_comp 0 eqs)end ));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   165
  fun add_one (thy,(dnam,axs,dfs)) = thy
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 23152
diff changeset
   166
	|> Sign.add_path dnam
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   167
	|> add_defs_infer dfs
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   168
	|> add_axioms_infer axs
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 23152
diff changeset
   169
	|> Sign.parent_path;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   170
  val thy = Library.foldl add_one (thy', mapn (calc_axioms comp_dname eqs) 0 eqs);
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 23152
diff changeset
   171
in thy |> Sign.add_path comp_dnam  
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   172
       |> add_defs_infer (bisim_def::(if length eqs>1 then [copy_def] else []))
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 23152
diff changeset
   173
       |> Sign.parent_path
30240
blanchet
parents: 29585
diff changeset
   174
       |> fold add_matchers eqs
blanchet
parents: 29585
diff changeset
   175
end; (* let (add_axioms) *)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   176
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   177
end; (* local *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   178
end; (* struct *)