| author | wenzelm | 
| Thu, 26 Aug 2010 11:29:43 +0200 | |
| changeset 38725 | 3d9d5ff80f6f | 
| parent 37216 | 3165bc303f66 | 
| child 45289 | 25e9e7f527b4 | 
| permissions | -rw-r--r-- | 
| 21476 | 1  | 
(* Title: Pure/morphism.ML  | 
2  | 
Author: Makarius  | 
|
3  | 
||
4  | 
Abstract morphisms on formal entities.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
infix 1 $>  | 
|
8  | 
||
9  | 
signature BASIC_MORPHISM =  | 
|
10  | 
sig  | 
|
11  | 
type morphism  | 
|
| 24031 | 12  | 
type declaration = morphism -> Context.generic -> Context.generic  | 
| 21476 | 13  | 
val $> : morphism * morphism -> morphism  | 
14  | 
end  | 
|
15  | 
||
16  | 
signature MORPHISM =  | 
|
17  | 
sig  | 
|
18  | 
include BASIC_MORPHISM  | 
|
| 29581 | 19  | 
val binding: morphism -> binding -> binding  | 
| 21476 | 20  | 
val typ: morphism -> typ -> typ  | 
21  | 
val term: morphism -> term -> term  | 
|
| 21521 | 22  | 
val fact: morphism -> thm list -> thm list  | 
| 21476 | 23  | 
val thm: morphism -> thm -> thm  | 
| 22235 | 24  | 
val cterm: morphism -> cterm -> cterm  | 
| 21476 | 25  | 
val morphism:  | 
| 29581 | 26  | 
   {binding: binding -> binding,
 | 
| 21476 | 27  | 
typ: typ -> typ,  | 
28  | 
term: term -> term,  | 
|
| 21521 | 29  | 
fact: thm list -> thm list} -> morphism  | 
| 29581 | 30  | 
val binding_morphism: (binding -> binding) -> morphism  | 
| 21492 | 31  | 
val typ_morphism: (typ -> typ) -> morphism  | 
32  | 
val term_morphism: (term -> term) -> morphism  | 
|
| 21521 | 33  | 
val fact_morphism: (thm list -> thm list) -> morphism  | 
| 21492 | 34  | 
val thm_morphism: (thm -> thm) -> morphism  | 
| 21476 | 35  | 
val identity: morphism  | 
| 
22571
 
3f00e937d1c9
renamed comp to compose (avoid clash with Alice keywords);
 
wenzelm 
parents: 
22235 
diff
changeset
 | 
36  | 
val compose: morphism -> morphism -> morphism  | 
| 
22670
 
c803b2696ada
added Morphism.transform/form (generic non-sense);
 
wenzelm 
parents: 
22571 
diff
changeset
 | 
37  | 
val transform: morphism -> (morphism -> 'a) -> morphism -> 'a  | 
| 
 
c803b2696ada
added Morphism.transform/form (generic non-sense);
 
wenzelm 
parents: 
22571 
diff
changeset
 | 
38  | 
val form: (morphism -> 'a) -> 'a  | 
| 21476 | 39  | 
end;  | 
40  | 
||
41  | 
structure Morphism: MORPHISM =  | 
|
42  | 
struct  | 
|
43  | 
||
44  | 
datatype morphism = Morphism of  | 
|
| 29581 | 45  | 
 {binding: binding -> binding,
 | 
| 21476 | 46  | 
typ: typ -> typ,  | 
47  | 
term: term -> term,  | 
|
| 21521 | 48  | 
fact: thm list -> thm list};  | 
| 21476 | 49  | 
|
| 24031 | 50  | 
type declaration = morphism -> Context.generic -> Context.generic;  | 
51  | 
||
| 28965 | 52  | 
fun binding (Morphism {binding, ...}) = binding;
 | 
| 21476 | 53  | 
fun typ (Morphism {typ, ...}) = typ;
 | 
54  | 
fun term (Morphism {term, ...}) = term;
 | 
|
| 21521 | 55  | 
fun fact (Morphism {fact, ...}) = fact;
 | 
56  | 
val thm = singleton o fact;  | 
|
| 22235 | 57  | 
val cterm = Drule.cterm_rule o thm;  | 
| 21476 | 58  | 
|
59  | 
val morphism = Morphism;  | 
|
60  | 
||
| 29605 | 61  | 
fun binding_morphism binding = morphism {binding = binding, typ = I, term = I, fact = I};
 | 
62  | 
fun typ_morphism typ = morphism {binding = I, typ = typ, term = I, fact = I};
 | 
|
63  | 
fun term_morphism term = morphism {binding = I, typ = I, term = term, fact = I};
 | 
|
64  | 
fun fact_morphism fact = morphism {binding = I, typ = I, term = I, fact = fact};
 | 
|
65  | 
fun thm_morphism thm = morphism {binding = I, typ = I, term = I, fact = map thm};
 | 
|
| 21492 | 66  | 
|
| 29605 | 67  | 
val identity = morphism {binding = I, typ = I, term = I, fact = I};
 | 
| 21492 | 68  | 
|
| 
22571
 
3f00e937d1c9
renamed comp to compose (avoid clash with Alice keywords);
 
wenzelm 
parents: 
22235 
diff
changeset
 | 
69  | 
fun compose  | 
| 29605 | 70  | 
    (Morphism {binding = binding1, typ = typ1, term = term1, fact = fact1})
 | 
71  | 
    (Morphism {binding = binding2, typ = typ2, term = term2, fact = fact2}) =
 | 
|
72  | 
  morphism {binding = binding1 o binding2, typ = typ1 o typ2,
 | 
|
73  | 
term = term1 o term2, fact = fact1 o fact2};  | 
|
| 21476 | 74  | 
|
| 
22571
 
3f00e937d1c9
renamed comp to compose (avoid clash with Alice keywords);
 
wenzelm 
parents: 
22235 
diff
changeset
 | 
75  | 
fun phi1 $> phi2 = compose phi2 phi1;  | 
| 21476 | 76  | 
|
| 
22670
 
c803b2696ada
added Morphism.transform/form (generic non-sense);
 
wenzelm 
parents: 
22571 
diff
changeset
 | 
77  | 
fun transform phi f = fn psi => f (phi $> psi);  | 
| 
 
c803b2696ada
added Morphism.transform/form (generic non-sense);
 
wenzelm 
parents: 
22571 
diff
changeset
 | 
78  | 
fun form f = f identity;  | 
| 
 
c803b2696ada
added Morphism.transform/form (generic non-sense);
 
wenzelm 
parents: 
22571 
diff
changeset
 | 
79  | 
|
| 21476 | 80  | 
end;  | 
81  | 
||
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
29605 
diff
changeset
 | 
82  | 
structure Basic_Morphism: BASIC_MORPHISM = Morphism;  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
29605 
diff
changeset
 | 
83  | 
open Basic_Morphism;  |