src/HOLCF/domain/axioms.ML
author wenzelm
Sat, 03 Nov 2001 01:41:26 +0100
changeset 12030 46d57d0290a2
parent 11531 d038246a62f2
child 14981 e73f8140af78
permissions -rw-r--r--
GPLed;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2453
2d416226b27d corrected headers
oheimb
parents: 1637
diff changeset
     1
(*  Title:      HOLCF/domain/axioms.ML
2d416226b27d corrected headers
oheimb
parents: 1637
diff changeset
     2
    ID:         $Id$
12030
wenzelm
parents: 11531
diff changeset
     3
    Author:     David von Oheimb
wenzelm
parents: 11531
diff changeset
     4
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
2453
2d416226b27d corrected headers
oheimb
parents: 1637
diff changeset
     5
12030
wenzelm
parents: 11531
diff changeset
     6
Syntax generator for domain section.
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
     7
*)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
     8
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
     9
structure Domain_Axioms = struct
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    10
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    11
local
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    12
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    13
open Domain_Library;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    14
infixr 0 ===>;infixr 0 ==>;infix 0 == ; 
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    15
infix 1 ===; infix 1 ~= ; infix 1 <<; infix 1 ~<<;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    16
infix 9 `   ; infix 9 `% ; infix 9 `%%; infixr 9 oo;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    17
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    18
fun infer_types thy' = map (inferT_axm (sign_of thy'));
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    19
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    20
fun calc_axioms comp_dname (eqs : eq list) n (((dname,_),cons) : eq)=
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    21
let
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    22
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    23
(* ----- axioms and definitions concerning the isomorphism ------------------ *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    24
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    25
  val dc_abs = %%:(dname^"_abs");
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    26
  val dc_rep = %%:(dname^"_rep");
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    27
  val x_name'= "x";
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    28
  val x_name = idx_name eqs x_name' (n+1);
4008
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3771
diff changeset
    29
  val dnam = Sign.base_name dname;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    30
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    31
 val abs_iso_ax = ("abs_iso" ,mk_trp(dc_rep`(dc_abs`%x_name')=== %:x_name'));
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    32
 val rep_iso_ax = ("rep_iso" ,mk_trp(dc_abs`(dc_rep`%x_name')=== %:x_name'));
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    33
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    34
  val when_def = ("when_def",%%:(dname^"_when") == 
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    35
     foldr (uncurry /\ ) (when_funs cons, /\x_name'((when_body cons (fn (x,y) =>
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    36
				Bound(1+length cons+x-y)))`(dc_rep`Bound 0))));
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    37
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    38
  fun con_def outer recu m n (_,args) = let
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    39
     fun idxs z x arg = (if is_lazy arg then fn t => %%:"up"`t else Id)
4755
843b5f159c7e added cproj', and therefore extended prj
oheimb
parents: 4122
diff changeset
    40
			(if recu andalso is_rec arg then (cproj (Bound z) eqs
843b5f159c7e added cproj', and therefore extended prj
oheimb
parents: 4122
diff changeset
    41
				  (rec_of arg))`Bound(z-x) else Bound(z-x));
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    42
     fun parms [] = %%:"ONE"
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    43
     |   parms vs = foldr'(fn(x,t)=> %%:"spair"`x`t)(mapn (idxs(length vs))1 vs);
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    44
     fun inj y 1 _ = y
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    45
     |   inj y _ 0 = %%:"sinl"`y
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    46
     |   inj y i j = %%:"sinr"`(inj y (i-1) (j-1));
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    47
  in foldr /\# (args, outer (inj (parms args) m n)) end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    48
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    49
  val copy_def = ("copy_def", %%:(dname^"_copy") == /\"f" (dc_abs oo 
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    50
	foldl (op `) (%%:(dname^"_when") , 
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    51
	              mapn (con_def Id true (length cons)) 0 cons)));
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    52
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    53
(* -- definitions concerning the constructors, discriminators and selectors - *)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    54
4043
35766855f344 domain package:
oheimb
parents: 4008
diff changeset
    55
  val con_defs = mapn (fn n => fn (con,args) => (extern_name con ^"_def",  
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    56
  %%:con == con_def (fn t => dc_abs`t) false (length cons) n (con,args))) 0 cons;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    57
4043
35766855f344 domain package:
oheimb
parents: 4008
diff changeset
    58
  val dis_defs = let
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    59
	fun ddef (con,_) = (dis_name con ^"_def",%%:(dis_name con) == 
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    60
		 mk_cRep_CFun(%%:(dname^"_when"),map 
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    61
			(fn (con',args) => (foldr /\#
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    62
			   (args,if con'=con then %%:"TT" else %%:"FF"))) cons))
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    63
	in map ddef cons end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    64
4043
35766855f344 domain package:
oheimb
parents: 4008
diff changeset
    65
  val sel_defs = let
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    66
	fun sdef con n arg = (sel_of arg^"_def",%%:(sel_of arg) == 
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    67
		 mk_cRep_CFun(%%:(dname^"_when"),map 
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    68
			(fn (con',args) => if con'<>con then %%:"UU" else
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    69
			 foldr /\# (args,Bound (length args - n))) cons));
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    70
	in flat(map (fn (con,args) => mapn (sdef con) 1 args) cons) end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    71
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    72
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    73
(* ----- axiom and definitions concerning induction ------------------------- *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    74
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    75
  val reach_ax = ("reach", mk_trp(cproj (%%:"fix"`%%(comp_dname^"_copy")) eqs n
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    76
					`%x_name === %:x_name));
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    77
  val take_def = ("take_def",%%:(dname^"_take") == mk_lam("n",cproj' 
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    78
	     (%%:"iterate" $ Bound 0 $ %%:(comp_dname^"_copy") $ %%:"UU") eqs n));
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    79
  val finite_def = ("finite_def",%%:(dname^"_finite") == mk_lam(x_name,
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    80
	mk_ex("n",(%%:(dname^"_take") $ Bound 0)`Bound 1 === Bound 1)));
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    81
4043
35766855f344 domain package:
oheimb
parents: 4008
diff changeset
    82
in (dnam,
35766855f344 domain package:
oheimb
parents: 4008
diff changeset
    83
    [abs_iso_ax, rep_iso_ax, reach_ax],
35766855f344 domain package:
oheimb
parents: 4008
diff changeset
    84
    [when_def, copy_def] @
35766855f344 domain package:
oheimb
parents: 4008
diff changeset
    85
     con_defs @ dis_defs @ sel_defs @
35766855f344 domain package:
oheimb
parents: 4008
diff changeset
    86
    [take_def, finite_def])
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    87
end; (* let *)
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    88
8438
b8389b4fca9c adapted to new PureThy.add_thms etc.;
wenzelm
parents: 6092
diff changeset
    89
fun add_axioms_i x = #1 o PureThy.add_axioms_i (map Thm.no_attributes x);
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    90
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    91
in (* local *)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    92
4008
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3771
diff changeset
    93
fun add_axioms (comp_dnam, eqs : eq list) thy' = let
2444085532c6 adapted domain and ax_ops package for name spaces
oheimb
parents: 3771
diff changeset
    94
  val comp_dname = Sign.full_name (sign_of thy') comp_dnam;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    95
  val dnames = map (fst o fst) eqs;
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
    96
  val x_name = idx_name dnames "x"; 
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    97
  fun copy_app dname = %%:(dname^"_copy")`Bound 0;
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
    98
  val copy_def = ("copy_def" , %%:(comp_dname^"_copy") ==
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
    99
				    /\"f"(foldr' cpair (map copy_app dnames)));
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
   100
  val bisim_def = ("bisim_def",%%:(comp_dname^"_bisim")==mk_lam("R",
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   101
    let
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   102
      fun one_con (con,args) = let
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   103
	val nonrec_args = filter_out is_rec args;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   104
	val    rec_args = filter     is_rec args;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   105
	val    recs_cnt = length rec_args;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   106
	val allargs     = nonrec_args @ rec_args
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   107
				      @ map (upd_vname (fn s=> s^"'")) rec_args;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   108
	val allvns      = map vname allargs;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   109
	fun vname_arg s arg = if is_rec arg then vname arg^s else vname arg;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   110
	val vns1        = map (vname_arg "" ) args;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   111
	val vns2        = map (vname_arg "'") args;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   112
	val allargs_cnt = length nonrec_args + 2*recs_cnt;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   113
	val rec_idxs    = (recs_cnt-1) downto 0;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   114
	val nonlazy_idxs = map snd (filter_out (fn (arg,_) => is_lazy arg)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   115
					 (allargs~~((allargs_cnt-1) downto 0)));
4755
843b5f159c7e added cproj', and therefore extended prj
oheimb
parents: 4122
diff changeset
   116
	fun rel_app i ra = proj (Bound(allargs_cnt+2)) eqs (rec_of ra) $ 
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   117
			   Bound (2*recs_cnt-i) $ Bound (recs_cnt-i);
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   118
	val capps = foldr mk_conj (mapn rel_app 1 rec_args, mk_conj(
11531
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
   119
	   Bound(allargs_cnt+1)===mk_cRep_CFun(%%:con,map (bound_arg allvns) vns1),
d038246a62f2 Renamed functions % and %% to avoid clash with syntax for proof terms.
berghofe
parents: 8438
diff changeset
   120
	   Bound(allargs_cnt+0)===mk_cRep_CFun(%%:con,map (bound_arg allvns) vns2)));
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   121
        in foldr mk_ex (allvns, foldr mk_conj 
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   122
			      (map (defined o Bound) nonlazy_idxs,capps)) end;
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   123
      fun one_comp n (_,cons) =mk_all(x_name(n+1),mk_all(x_name(n+1)^"'",mk_imp(
4755
843b5f159c7e added cproj', and therefore extended prj
oheimb
parents: 4122
diff changeset
   124
	 		proj (Bound 2) eqs n $ Bound 1 $ Bound 0,
1637
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   125
         		foldr' mk_disj (mk_conj(Bound 1 === UU,Bound 0 === UU)
b8a8ae2e5de1 Updated: 01-Mar-96 when functional strictified, copy_def based on when_def
oheimb
parents: 1461
diff changeset
   126
					::map one_con cons))));
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   127
    in foldr' mk_conj (mapn one_comp 0 eqs)end ));
4043
35766855f344 domain package:
oheimb
parents: 4008
diff changeset
   128
  fun add_one (thy,(dnam,axs,dfs)) = thy 
35766855f344 domain package:
oheimb
parents: 4008
diff changeset
   129
	|> Theory.add_path dnam
4862
613ce4cc303f adapted to new PureThy.add_axioms_i;
wenzelm
parents: 4755
diff changeset
   130
	|> add_axioms_i (infer_types thy' dfs)(*add_defs_i*)
613ce4cc303f adapted to new PureThy.add_axioms_i;
wenzelm
parents: 4755
diff changeset
   131
	|> add_axioms_i (infer_types thy' axs)
613ce4cc303f adapted to new PureThy.add_axioms_i;
wenzelm
parents: 4755
diff changeset
   132
	|> Theory.parent_path;
4043
35766855f344 domain package:
oheimb
parents: 4008
diff changeset
   133
  val thy = foldl add_one (thy', mapn (calc_axioms comp_dname eqs) 0 eqs);
35766855f344 domain package:
oheimb
parents: 4008
diff changeset
   134
in thy |> Theory.add_path comp_dnam  
4862
613ce4cc303f adapted to new PureThy.add_axioms_i;
wenzelm
parents: 4755
diff changeset
   135
       |> add_axioms_i (infer_types thy' 
4043
35766855f344 domain package:
oheimb
parents: 4008
diff changeset
   136
		(bisim_def::(if length eqs>1 then [copy_def] else [])))
4862
613ce4cc303f adapted to new PureThy.add_axioms_i;
wenzelm
parents: 4755
diff changeset
   137
       |> Theory.parent_path
4043
35766855f344 domain package:
oheimb
parents: 4008
diff changeset
   138
end;
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   139
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   140
end; (* local *)
ea0668a1c0ba added 8bit pragmas
regensbu
parents:
diff changeset
   141
end; (* struct *)