| author | kuncar | 
| Wed, 05 Jun 2013 15:21:52 +0200 | |
| changeset 52307 | 32c433c38ddd | 
| parent 45289 | 25e9e7f527b4 | 
| child 53087 | 5a1dcda7967c | 
| 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  | 
|
| 
45289
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
19  | 
  type 'a funs = ('a -> 'a) list
 | 
| 29581 | 20  | 
val binding: morphism -> binding -> binding  | 
| 21476 | 21  | 
val typ: morphism -> typ -> typ  | 
22  | 
val term: morphism -> term -> term  | 
|
| 21521 | 23  | 
val fact: morphism -> thm list -> thm list  | 
| 21476 | 24  | 
val thm: morphism -> thm -> thm  | 
| 22235 | 25  | 
val cterm: morphism -> cterm -> cterm  | 
| 21476 | 26  | 
val morphism:  | 
| 
45289
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
27  | 
   {binding: binding funs,
 | 
| 
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
28  | 
typ: typ funs,  | 
| 
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
29  | 
term: term funs,  | 
| 
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
30  | 
fact: thm list funs} -> morphism  | 
| 29581 | 31  | 
val binding_morphism: (binding -> binding) -> morphism  | 
| 21492 | 32  | 
val typ_morphism: (typ -> typ) -> morphism  | 
33  | 
val term_morphism: (term -> term) -> morphism  | 
|
| 21521 | 34  | 
val fact_morphism: (thm list -> thm list) -> morphism  | 
| 21492 | 35  | 
val thm_morphism: (thm -> thm) -> morphism  | 
| 21476 | 36  | 
val identity: morphism  | 
| 
22571
 
3f00e937d1c9
renamed comp to compose (avoid clash with Alice keywords);
 
wenzelm 
parents: 
22235 
diff
changeset
 | 
37  | 
val compose: morphism -> morphism -> morphism  | 
| 
22670
 
c803b2696ada
added Morphism.transform/form (generic non-sense);
 
wenzelm 
parents: 
22571 
diff
changeset
 | 
38  | 
val transform: morphism -> (morphism -> 'a) -> morphism -> 'a  | 
| 
 
c803b2696ada
added Morphism.transform/form (generic non-sense);
 
wenzelm 
parents: 
22571 
diff
changeset
 | 
39  | 
val form: (morphism -> 'a) -> 'a  | 
| 21476 | 40  | 
end;  | 
41  | 
||
42  | 
structure Morphism: MORPHISM =  | 
|
43  | 
struct  | 
|
44  | 
||
| 
45289
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
45  | 
type 'a funs = ('a -> 'a) list;
 | 
| 
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
46  | 
fun apply fs = fold_rev (fn f => fn x => f x) fs;  | 
| 
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
47  | 
|
| 21476 | 48  | 
datatype morphism = Morphism of  | 
| 
45289
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
49  | 
 {binding: binding funs,
 | 
| 
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
50  | 
typ: typ funs,  | 
| 
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
51  | 
term: term funs,  | 
| 
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
52  | 
fact: thm list funs};  | 
| 21476 | 53  | 
|
| 24031 | 54  | 
type declaration = morphism -> Context.generic -> Context.generic;  | 
55  | 
||
| 
45289
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
56  | 
fun binding (Morphism {binding, ...}) = apply binding;
 | 
| 
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
57  | 
fun typ (Morphism {typ, ...}) = apply typ;
 | 
| 
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
58  | 
fun term (Morphism {term, ...}) = apply term;
 | 
| 
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
59  | 
fun fact (Morphism {fact, ...}) = apply fact;
 | 
| 21521 | 60  | 
val thm = singleton o fact;  | 
| 22235 | 61  | 
val cterm = Drule.cterm_rule o thm;  | 
| 21476 | 62  | 
|
63  | 
val morphism = Morphism;  | 
|
64  | 
||
| 
45289
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
65  | 
fun binding_morphism binding = morphism {binding = [binding], typ = [], term = [], fact = []};
 | 
| 
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
66  | 
fun typ_morphism typ = morphism {binding = [], typ = [typ], term = [], fact = []};
 | 
| 
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
67  | 
fun term_morphism term = morphism {binding = [], typ = [], term = [term], fact = []};
 | 
| 
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
68  | 
fun fact_morphism fact = morphism {binding = [], typ = [], term = [], fact = [fact]};
 | 
| 
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
69  | 
fun thm_morphism thm = morphism {binding = [], typ = [], term = [], fact = [map thm]};
 | 
| 21492 | 70  | 
|
| 
45289
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
71  | 
val identity = morphism {binding = [], typ = [], term = [], fact = []};
 | 
| 21492 | 72  | 
|
| 
22571
 
3f00e937d1c9
renamed comp to compose (avoid clash with Alice keywords);
 
wenzelm 
parents: 
22235 
diff
changeset
 | 
73  | 
fun compose  | 
| 29605 | 74  | 
    (Morphism {binding = binding1, typ = typ1, term = term1, fact = fact1})
 | 
75  | 
    (Morphism {binding = binding2, typ = typ2, term = term2, fact = fact2}) =
 | 
|
| 
45289
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
76  | 
  morphism {binding = binding1 @ binding2, typ = typ1 @ typ2,
 | 
| 
 
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
77  | 
term = term1 @ term2, fact = fact1 @ fact2};  | 
| 21476 | 78  | 
|
| 
22571
 
3f00e937d1c9
renamed comp to compose (avoid clash with Alice keywords);
 
wenzelm 
parents: 
22235 
diff
changeset
 | 
79  | 
fun phi1 $> phi2 = compose phi2 phi1;  | 
| 21476 | 80  | 
|
| 
22670
 
c803b2696ada
added Morphism.transform/form (generic non-sense);
 
wenzelm 
parents: 
22571 
diff
changeset
 | 
81  | 
fun transform phi f = fn psi => f (phi $> psi);  | 
| 
 
c803b2696ada
added Morphism.transform/form (generic non-sense);
 
wenzelm 
parents: 
22571 
diff
changeset
 | 
82  | 
fun form f = f identity;  | 
| 
 
c803b2696ada
added Morphism.transform/form (generic non-sense);
 
wenzelm 
parents: 
22571 
diff
changeset
 | 
83  | 
|
| 21476 | 84  | 
end;  | 
85  | 
||
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
29605 
diff
changeset
 | 
86  | 
structure Basic_Morphism: BASIC_MORPHISM = Morphism;  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
29605 
diff
changeset
 | 
87  | 
open Basic_Morphism;  |