| author | wenzelm | 
| Sun, 30 Apr 2017 17:02:56 +0200 | |
| changeset 65647 | 7cf60e2b9115 | 
| parent 63568 | e63c8f2fbd28 | 
| child 67149 | e61557884799 | 
| permissions | -rw-r--r-- | 
| 55467 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 blanchet parents: 
51717diff
changeset | 1 | (* Title: HOL/Tools/functor.ML | 
| 40582 | 2 | Author: Florian Haftmann, TU Muenchen | 
| 3 | ||
| 40968 
a6fcd305f7dc
replace `type_mapper` by the more adequate `type_lifting`
 haftmann parents: 
40857diff
changeset | 4 | Functorial structure of types. | 
| 40582 | 5 | *) | 
| 6 | ||
| 55467 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 blanchet parents: 
51717diff
changeset | 7 | signature FUNCTOR = | 
| 40582 | 8 | sig | 
| 41388 | 9 | val find_atomic: Proof.context -> typ -> (typ * (bool * bool)) list | 
| 10 | val construct_mapper: Proof.context -> (string * bool -> term) | |
| 40582 | 11 | -> bool -> typ -> typ -> term | 
| 60488 | 12 | val functor_: string option -> term -> local_theory -> Proof.state | 
| 55467 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 blanchet parents: 
51717diff
changeset | 13 | val functor_cmd: string option -> string -> Proof.context -> Proof.state | 
| 40582 | 14 | type entry | 
| 41390 
207ee8f8a19c
full localization with possibly multiple entries;
 haftmann parents: 
41389diff
changeset | 15 | val entries: Proof.context -> entry list Symtab.table | 
| 40582 | 16 | end; | 
| 17 | ||
| 55467 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 blanchet parents: 
51717diff
changeset | 18 | structure Functor : FUNCTOR = | 
| 40582 | 19 | struct | 
| 20 | ||
| 41395 | 21 | (* bookkeeping *) | 
| 22 | ||
| 41371 
35d2241c169c
prove more algebraic version of functorial properties; retain old properties for convenience
 haftmann parents: 
41298diff
changeset | 23 | val compN = "comp"; | 
| 
35d2241c169c
prove more algebraic version of functorial properties; retain old properties for convenience
 haftmann parents: 
41298diff
changeset | 24 | val idN = "id"; | 
| 40611 | 25 | val compositionalityN = "compositionality"; | 
| 40594 
fae1da97bb5e
infer variances of user-given mapper operation; proper thm storing
 haftmann parents: 
40587diff
changeset | 26 | val identityN = "identity"; | 
| 
fae1da97bb5e
infer variances of user-given mapper operation; proper thm storing
 haftmann parents: 
40587diff
changeset | 27 | |
| 41387 | 28 | type entry = { mapper: term, variances: (sort * (bool * bool)) list,
 | 
| 41371 
35d2241c169c
prove more algebraic version of functorial properties; retain old properties for convenience
 haftmann parents: 
41298diff
changeset | 29 | comp: thm, id: thm }; | 
| 40582 | 30 | |
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
41395diff
changeset | 31 | structure Data = Generic_Data | 
| 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
41395diff
changeset | 32 | ( | 
| 41390 
207ee8f8a19c
full localization with possibly multiple entries;
 haftmann parents: 
41389diff
changeset | 33 | type T = entry list Symtab.table | 
| 40582 | 34 | val empty = Symtab.empty | 
| 35 | val extend = I | |
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
41395diff
changeset | 36 | fun merge data = Symtab.merge (K true) data | 
| 40582 | 37 | ); | 
| 38 | ||
| 41388 | 39 | val entries = Data.get o Context.Proof; | 
| 40582 | 40 | |
| 41 | ||
| 42 | (* type analysis *) | |
| 43 | ||
| 59838 | 44 | fun term_with_typ ctxt T t = | 
| 45 | Envir.subst_term_types | |
| 46 | (Sign.typ_match (Proof_Context.theory_of ctxt) (fastype_of t, T) Vartab.empty) t; | |
| 41389 | 47 | |
| 41388 | 48 | fun find_atomic ctxt T = | 
| 40582 | 49 | let | 
| 41390 
207ee8f8a19c
full localization with possibly multiple entries;
 haftmann parents: 
41389diff
changeset | 50 | val variances_of = Option.map #variances o try hd o Symtab.lookup_list (entries ctxt); | 
| 40582 | 51 | fun add_variance is_contra T = | 
| 52 | AList.map_default (op =) (T, (false, false)) | |
| 53 | ((if is_contra then apsnd else apfst) (K true)); | |
| 54 | fun analyze' is_contra (_, (co, contra)) T = | |
| 55 | (if co then analyze is_contra T else I) | |
| 56 | #> (if contra then analyze (not is_contra) T else I) | |
| 57 | and analyze is_contra (T as Type (tyco, Ts)) = (case variances_of tyco | |
| 58 | of NONE => add_variance is_contra T | |
| 59 | | SOME variances => fold2 (analyze' is_contra) variances Ts) | |
| 60 | | analyze is_contra T = add_variance is_contra T; | |
| 61 | in analyze false T [] end; | |
| 62 | ||
| 41388 | 63 | fun construct_mapper ctxt atomic = | 
| 40582 | 64 | let | 
| 41390 
207ee8f8a19c
full localization with possibly multiple entries;
 haftmann parents: 
41389diff
changeset | 65 | val lookup = hd o Symtab.lookup_list (entries ctxt); | 
| 40582 | 66 | fun constructs is_contra (_, (co, contra)) T T' = | 
| 67 | (if co then [construct is_contra T T'] else []) | |
| 68 | @ (if contra then [construct (not is_contra) T T'] else []) | |
| 69 | and construct is_contra (T as Type (tyco, Ts)) (T' as Type (_, Ts')) = | |
| 70 | let | |
| 41388 | 71 |             val { mapper = raw_mapper, variances, ... } = lookup tyco;
 | 
| 40582 | 72 | val args = maps (fn (arg_pattern, (T, T')) => | 
| 73 | constructs is_contra arg_pattern T T') | |
| 74 | (variances ~~ (Ts ~~ Ts')); | |
| 75 | val (U, U') = if is_contra then (T', T) else (T, T'); | |
| 41388 | 76 | val mapper = term_with_typ ctxt (map fastype_of args ---> U --> U') raw_mapper; | 
| 77 | in list_comb (mapper, args) end | |
| 40582 | 78 | | construct is_contra (TFree (v, _)) (TFree _) = atomic (v, is_contra); | 
| 79 | in construct end; | |
| 80 | ||
| 81 | ||
| 82 | (* mapper properties *) | |
| 83 | ||
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51551diff
changeset | 84 | val compositionality_ss = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51551diff
changeset | 85 |   simpset_of (put_simpset HOL_basic_ss @{context} addsimps [Simpdata.mk_eq @{thm comp_def}]);
 | 
| 41387 | 86 | |
| 41371 
35d2241c169c
prove more algebraic version of functorial properties; retain old properties for convenience
 haftmann parents: 
41298diff
changeset | 87 | fun make_comp_prop ctxt variances (tyco, mapper) = | 
| 40582 | 88 | let | 
| 41371 
35d2241c169c
prove more algebraic version of functorial properties; retain old properties for convenience
 haftmann parents: 
41298diff
changeset | 89 | val sorts = map fst variances | 
| 
35d2241c169c
prove more algebraic version of functorial properties; retain old properties for convenience
 haftmann parents: 
41298diff
changeset | 90 | val (((vs3, vs2), vs1), _) = ctxt | 
| 
35d2241c169c
prove more algebraic version of functorial properties; retain old properties for convenience
 haftmann parents: 
41298diff
changeset | 91 | |> Variable.invent_types sorts | 
| 
35d2241c169c
prove more algebraic version of functorial properties; retain old properties for convenience
 haftmann parents: 
41298diff
changeset | 92 | ||>> Variable.invent_types sorts | 
| 
35d2241c169c
prove more algebraic version of functorial properties; retain old properties for convenience
 haftmann parents: 
41298diff
changeset | 93 | ||>> Variable.invent_types sorts | 
| 
35d2241c169c
prove more algebraic version of functorial properties; retain old properties for convenience
 haftmann parents: 
41298diff
changeset | 94 | val (Ts1, Ts2, Ts3) = (map TFree vs1, map TFree vs2, map TFree vs3); | 
| 40582 | 95 | fun mk_argT ((T, T'), (_, (co, contra))) = | 
| 96 | (if co then [(T --> T')] else []) | |
| 97 | @ (if contra then [(T' --> T)] else []); | |
| 98 | val contras = maps (fn (_, (co, contra)) => | |
| 99 | (if co then [false] else []) @ (if contra then [true] else [])) variances; | |
| 100 | val Ts21 = maps mk_argT ((Ts2 ~~ Ts1) ~~ variances); | |
| 101 | val Ts32 = maps mk_argT ((Ts3 ~~ Ts2) ~~ variances); | |
| 41371 
35d2241c169c
prove more algebraic version of functorial properties; retain old properties for convenience
 haftmann parents: 
41298diff
changeset | 102 | fun invents n k nctxt = | 
| 
35d2241c169c
prove more algebraic version of functorial properties; retain old properties for convenience
 haftmann parents: 
41298diff
changeset | 103 | let | 
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
42388diff
changeset | 104 | val names = Name.invent nctxt n k; | 
| 41371 
35d2241c169c
prove more algebraic version of functorial properties; retain old properties for convenience
 haftmann parents: 
41298diff
changeset | 105 | in (names, fold Name.declare names nctxt) end; | 
| 
35d2241c169c
prove more algebraic version of functorial properties; retain old properties for convenience
 haftmann parents: 
41298diff
changeset | 106 | val ((names21, names32), nctxt) = Variable.names_of ctxt | 
| 40582 | 107 | |> invents "f" (length Ts21) | 
| 108 | ||>> invents "f" (length Ts32); | |
| 109 | val T1 = Type (tyco, Ts1); | |
| 110 | val T2 = Type (tyco, Ts2); | |
| 111 | val T3 = Type (tyco, Ts3); | |
| 112 | val (args21, args32) = (names21 ~~ Ts21, names32 ~~ Ts32); | |
| 113 | val args31 = map2 (fn is_contra => fn ((f21, T21), (f32, T32)) => | |
| 114 | if not is_contra then | |
| 41371 
35d2241c169c
prove more algebraic version of functorial properties; retain old properties for convenience
 haftmann parents: 
41298diff
changeset | 115 | HOLogic.mk_comp (Free (f21, T21), Free (f32, T32)) | 
| 40582 | 116 | else | 
| 41371 
35d2241c169c
prove more algebraic version of functorial properties; retain old properties for convenience
 haftmann parents: 
41298diff
changeset | 117 | HOLogic.mk_comp (Free (f32, T32), Free (f21, T21)) | 
| 40582 | 118 | ) contras (args21 ~~ args32) | 
| 41395 | 119 | fun mk_mapper T T' args = list_comb | 
| 120 | (term_with_typ ctxt (map fastype_of args ---> T --> T') mapper, args); | |
| 41387 | 121 | val mapper21 = mk_mapper T2 T1 (map Free args21); | 
| 122 | val mapper32 = mk_mapper T3 T2 (map Free args32); | |
| 123 | val mapper31 = mk_mapper T3 T1 args31; | |
| 41395 | 124 | val eq1 = (HOLogic.mk_Trueprop o HOLogic.mk_eq) | 
| 125 | (HOLogic.mk_comp (mapper21, mapper32), mapper31); | |
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
42388diff
changeset | 126 | val x = Free (the_single (Name.invent nctxt (Long_Name.base_name tyco) 1), T3) | 
| 41395 | 127 | val eq2 = (HOLogic.mk_Trueprop o HOLogic.mk_eq) | 
| 128 | (mapper21 $ (mapper32 $ x), mapper31 $ x); | |
| 41387 | 129 | val comp_prop = fold_rev Logic.all (map Free (args21 @ args32)) eq1; | 
| 130 | val compositionality_prop = fold_rev Logic.all (map Free (args21 @ args32) @ [x]) eq2; | |
| 61841 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
60488diff
changeset | 131 | fun prove_compositionality ctxt comp_thm = | 
| 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
60488diff
changeset | 132 | Goal.prove_sorry ctxt [] [] compositionality_prop | 
| 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
60488diff
changeset | 133 |         (K (ALLGOALS (Method.insert_tac ctxt [@{thm fun_cong} OF [comp_thm]]
 | 
| 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
60488diff
changeset | 134 | THEN' Simplifier.asm_lr_simp_tac (put_simpset compositionality_ss ctxt) | 
| 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
60488diff
changeset | 135 | THEN_ALL_NEW (Goal.assume_rule_tac ctxt)))); | 
| 41387 | 136 | in (comp_prop, prove_compositionality) end; | 
| 137 | ||
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51551diff
changeset | 138 | val identity_ss = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51551diff
changeset | 139 |   simpset_of (put_simpset HOL_basic_ss @{context} addsimps [Simpdata.mk_eq @{thm id_def}]);
 | 
| 40582 | 140 | |
| 41371 
35d2241c169c
prove more algebraic version of functorial properties; retain old properties for convenience
 haftmann parents: 
41298diff
changeset | 141 | fun make_id_prop ctxt variances (tyco, mapper) = | 
| 40582 | 142 | let | 
| 46810 | 143 | val (vs, _) = Variable.invent_types (map fst variances) ctxt; | 
| 41371 
35d2241c169c
prove more algebraic version of functorial properties; retain old properties for convenience
 haftmann parents: 
41298diff
changeset | 144 | val Ts = map TFree vs; | 
| 40582 | 145 | fun bool_num b = if b then 1 else 0; | 
| 146 | fun mk_argT (T, (_, (co, contra))) = | |
| 41387 | 147 | replicate (bool_num co + bool_num contra) T | 
| 148 | val arg_Ts = maps mk_argT (Ts ~~ variances) | |
| 40582 | 149 | val T = Type (tyco, Ts); | 
| 41387 | 150 | val head = term_with_typ ctxt (map (fn T => T --> T) arg_Ts ---> T --> T) mapper; | 
| 151 | val lhs1 = list_comb (head, map (HOLogic.id_const) arg_Ts); | |
| 152 |     val lhs2 = list_comb (head, map (fn arg_T => Abs ("x", arg_T, Bound 0)) arg_Ts);
 | |
| 153 | val rhs = HOLogic.id_const T; | |
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
55467diff
changeset | 154 | val (id_prop, identity_prop) = | 
| 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
55467diff
changeset | 155 | apply2 (HOLogic.mk_Trueprop o HOLogic.mk_eq o rpair rhs) (lhs1, lhs2); | 
| 61841 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
60488diff
changeset | 156 | fun prove_identity ctxt id_thm = | 
| 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
60488diff
changeset | 157 | Goal.prove_sorry ctxt [] [] identity_prop | 
| 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
60488diff
changeset | 158 | (K (ALLGOALS (Method.insert_tac ctxt [id_thm] THEN' | 
| 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
60488diff
changeset | 159 | Simplifier.asm_lr_simp_tac (put_simpset identity_ss ctxt)))); | 
| 41387 | 160 | in (id_prop, prove_identity) end; | 
| 40582 | 161 | |
| 162 | ||
| 40597 | 163 | (* analyzing and registering mappers *) | 
| 40582 | 164 | |
| 55467 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 blanchet parents: 
51717diff
changeset | 165 | fun consume _ _ [] = (false, []) | 
| 40594 
fae1da97bb5e
infer variances of user-given mapper operation; proper thm storing
 haftmann parents: 
40587diff
changeset | 166 | | consume eq x (ys as z :: zs) = if eq (x, z) then (true, zs) else (false, ys); | 
| 
fae1da97bb5e
infer variances of user-given mapper operation; proper thm storing
 haftmann parents: 
40587diff
changeset | 167 | |
| 40587 | 168 | fun split_mapper_typ "fun" T = | 
| 169 | let | |
| 170 | val (Ts', T') = strip_type T; | |
| 171 | val (Ts'', T'') = split_last Ts'; | |
| 172 | val (Ts''', T''') = split_last Ts''; | |
| 173 | in (Ts''', T''', T'' --> T') end | |
| 46810 | 174 | | split_mapper_typ _ T = | 
| 40587 | 175 | let | 
| 176 | val (Ts', T') = strip_type T; | |
| 177 | val (Ts'', T'') = split_last Ts'; | |
| 178 | in (Ts'', T'', T') end; | |
| 179 | ||
| 46852 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 180 | fun analyze_mapper ctxt input_mapper = | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 181 | let | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 182 | val T = fastype_of input_mapper; | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 183 | val _ = Type.no_tvars T; | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 184 | val _ = | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 185 | if null (subtract (op =) (Term.add_tfreesT T []) (Term.add_tfrees input_mapper [])) | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 186 | then () | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 187 |       else error ("Illegal additional type variable(s) in term: " ^ Syntax.string_of_term ctxt input_mapper);
 | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 188 | val _ = | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 189 | if null (Term.add_vars (singleton | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 190 | (Variable.export_terms (Variable.auto_fixes input_mapper ctxt) ctxt) input_mapper) []) | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 191 | then () | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 192 |       else error ("Illegal locally free variable(s) in term: "
 | 
| 63568 | 193 | ^ Syntax.string_of_term ctxt input_mapper); | 
| 46852 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 194 | val mapper = singleton (Variable.polymorphic ctxt) input_mapper; | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 195 | val _ = | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 196 | if null (Term.add_tfreesT (fastype_of mapper) []) then () | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 197 |       else error ("Illegal locally fixed type variable(s) in type: " ^ Syntax.string_of_typ ctxt T);
 | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 198 | fun add_tycos (Type (tyco, Ts)) = insert (op =) tyco #> fold add_tycos Ts | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 199 | | add_tycos _ = I; | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 200 | val tycos = add_tycos T []; | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 201 | val tyco = if tycos = ["fun"] then "fun" | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 202 | else case remove (op =) "fun" tycos | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 203 | of [tyco] => tyco | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 204 |         | _ => error ("Bad number of type constructors: " ^ Syntax.string_of_typ ctxt T);
 | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 205 | in (mapper, T, tyco) end; | 
| 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 206 | |
| 41390 
207ee8f8a19c
full localization with possibly multiple entries;
 haftmann parents: 
41389diff
changeset | 207 | fun analyze_variances ctxt tyco T = | 
| 40587 | 208 | let | 
| 41390 
207ee8f8a19c
full localization with possibly multiple entries;
 haftmann parents: 
41389diff
changeset | 209 |     fun bad_typ () = error ("Bad mapper type: " ^ Syntax.string_of_typ ctxt T);
 | 
| 40587 | 210 | val (Ts, T1, T2) = split_mapper_typ tyco T | 
| 211 | handle List.Empty => bad_typ (); | |
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
55467diff
changeset | 212 | val _ = | 
| 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
55467diff
changeset | 213 | apply2 ((fn tyco' => if tyco' = tyco then () else bad_typ ()) o fst o dest_Type) (T1, T2) | 
| 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
55467diff
changeset | 214 | handle TYPE _ => bad_typ (); | 
| 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
55467diff
changeset | 215 | val (vs1, vs2) = | 
| 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
55467diff
changeset | 216 | apply2 (map dest_TFree o snd o dest_Type) (T1, T2) | 
| 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
55467diff
changeset | 217 | handle TYPE _ => bad_typ (); | 
| 40587 | 218 | val _ = if has_duplicates (eq_fst (op =)) (vs1 @ vs2) | 
| 219 | then bad_typ () else (); | |
| 46810 | 220 | fun check_variance_pair (var1 as (_, sort1), var2 as (_, sort2)) = | 
| 40594 
fae1da97bb5e
infer variances of user-given mapper operation; proper thm storing
 haftmann parents: 
40587diff
changeset | 221 | let | 
| 
fae1da97bb5e
infer variances of user-given mapper operation; proper thm storing
 haftmann parents: 
40587diff
changeset | 222 | val coT = TFree var1 --> TFree var2; | 
| 
fae1da97bb5e
infer variances of user-given mapper operation; proper thm storing
 haftmann parents: 
40587diff
changeset | 223 | val contraT = TFree var2 --> TFree var1; | 
| 42361 | 224 | val sort = Sign.inter_sort (Proof_Context.theory_of ctxt) (sort1, sort2); | 
| 40594 
fae1da97bb5e
infer variances of user-given mapper operation; proper thm storing
 haftmann parents: 
40587diff
changeset | 225 | in | 
| 
fae1da97bb5e
infer variances of user-given mapper operation; proper thm storing
 haftmann parents: 
40587diff
changeset | 226 | consume (op =) coT | 
| 
fae1da97bb5e
infer variances of user-given mapper operation; proper thm storing
 haftmann parents: 
40587diff
changeset | 227 | ##>> consume (op =) contraT | 
| 
fae1da97bb5e
infer variances of user-given mapper operation; proper thm storing
 haftmann parents: 
40587diff
changeset | 228 | #>> pair sort | 
| 
fae1da97bb5e
infer variances of user-given mapper operation; proper thm storing
 haftmann parents: 
40587diff
changeset | 229 | end; | 
| 
fae1da97bb5e
infer variances of user-given mapper operation; proper thm storing
 haftmann parents: 
40587diff
changeset | 230 | val (variances, left_variances) = fold_map check_variance_pair (vs1 ~~ vs2) Ts; | 
| 
fae1da97bb5e
infer variances of user-given mapper operation; proper thm storing
 haftmann parents: 
40587diff
changeset | 231 | val _ = if null left_variances then () else bad_typ (); | 
| 
fae1da97bb5e
infer variances of user-given mapper operation; proper thm storing
 haftmann parents: 
40587diff
changeset | 232 | in variances end; | 
| 40587 | 233 | |
| 55467 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 blanchet parents: 
51717diff
changeset | 234 | fun gen_functor prep_term some_prfx raw_mapper lthy = | 
| 40583 | 235 | let | 
| 46852 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 236 | val (mapper, T, tyco) = analyze_mapper lthy (prep_term lthy raw_mapper); | 
| 41371 
35d2241c169c
prove more algebraic version of functorial properties; retain old properties for convenience
 haftmann parents: 
41298diff
changeset | 237 | val prfx = the_default (Long_Name.base_name tyco) some_prfx; | 
| 41390 
207ee8f8a19c
full localization with possibly multiple entries;
 haftmann parents: 
41389diff
changeset | 238 | val variances = analyze_variances lthy tyco T; | 
| 
207ee8f8a19c
full localization with possibly multiple entries;
 haftmann parents: 
41389diff
changeset | 239 | val (comp_prop, prove_compositionality) = make_comp_prop lthy variances (tyco, mapper); | 
| 
207ee8f8a19c
full localization with possibly multiple entries;
 haftmann parents: 
41389diff
changeset | 240 | val (id_prop, prove_identity) = make_id_prop lthy variances (tyco, mapper); | 
| 40856 
3ad8a5925ba4
optional explicit prefix for type mapper theorems
 haftmann parents: 
40855diff
changeset | 241 | val qualify = Binding.qualify true prfx o Binding.name; | 
| 41389 | 242 | fun mapper_declaration comp_thm id_thm phi context = | 
| 243 | let | |
| 42388 | 244 | val typ_instance = Sign.typ_instance (Context.theory_of context); | 
| 41389 | 245 | val mapper' = Morphism.term phi mapper; | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
55467diff
changeset | 246 | val T_T' = apply2 fastype_of (mapper, mapper'); | 
| 46852 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 247 | val vars = Term.add_vars mapper' []; | 
| 42388 | 248 | in | 
| 46852 
0b8dd4c8c79a
more precise checking for wellformedness of mapper, before and after morphism application
 haftmann parents: 
46851diff
changeset | 249 | if null vars andalso typ_instance T_T' andalso typ_instance (swap T_T') | 
| 41390 
207ee8f8a19c
full localization with possibly multiple entries;
 haftmann parents: 
41389diff
changeset | 250 | then (Data.map o Symtab.cons_list) (tyco, | 
| 41389 | 251 |           { mapper = mapper', variances = variances,
 | 
| 41390 
207ee8f8a19c
full localization with possibly multiple entries;
 haftmann parents: 
41389diff
changeset | 252 | comp = Morphism.thm phi comp_thm, id = Morphism.thm phi id_thm }) context | 
| 41389 | 253 | else context | 
| 254 | end; | |
| 41387 | 255 | fun after_qed [single_comp_thm, single_id_thm] lthy = | 
| 40587 | 256 | lthy | 
| 41387 | 257 | |> Local_Theory.note ((qualify compN, []), single_comp_thm) | 
| 258 | ||>> Local_Theory.note ((qualify idN, []), single_id_thm) | |
| 259 | |-> (fn ((_, [comp_thm]), (_, [id_thm])) => fn lthy => | |
| 41371 
35d2241c169c
prove more algebraic version of functorial properties; retain old properties for convenience
 haftmann parents: 
41298diff
changeset | 260 | lthy | 
| 41388 | 261 | |> Local_Theory.note ((qualify compositionalityN, []), | 
| 262 | [prove_compositionality lthy comp_thm]) | |
| 41371 
35d2241c169c
prove more algebraic version of functorial properties; retain old properties for convenience
 haftmann parents: 
41298diff
changeset | 263 | |> snd | 
| 41388 | 264 | |> Local_Theory.note ((qualify identityN, []), | 
| 265 | [prove_identity lthy id_thm]) | |
| 266 | |> snd | |
| 45291 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 wenzelm parents: 
43329diff
changeset | 267 |         |> Local_Theory.declaration {syntax = false, pervasive = false}
 | 
| 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 wenzelm parents: 
43329diff
changeset | 268 | (mapper_declaration comp_thm id_thm)) | 
| 40583 | 269 | in | 
| 41390 
207ee8f8a19c
full localization with possibly multiple entries;
 haftmann parents: 
41389diff
changeset | 270 | lthy | 
| 41371 
35d2241c169c
prove more algebraic version of functorial properties; retain old properties for convenience
 haftmann parents: 
41298diff
changeset | 271 | |> Proof.theorem NONE after_qed (map (fn t => [(t, [])]) [comp_prop, id_prop]) | 
| 40583 | 272 | end | 
| 273 | ||
| 60488 | 274 | val functor_ = gen_functor Syntax.check_term; | 
| 55467 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 blanchet parents: 
51717diff
changeset | 275 | val functor_cmd = gen_functor Syntax.read_term; | 
| 40583 | 276 | |
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46949diff
changeset | 277 | val _ = | 
| 59936 
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
 wenzelm parents: 
59838diff
changeset | 278 |   Outer_Syntax.local_theory_to_proof @{command_keyword functor}
 | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46949diff
changeset | 279 | "register operations managing the functorial structure of a type" | 
| 63120 
629a4c5e953e
embedded content may be delimited via cartouches;
 wenzelm parents: 
61841diff
changeset | 280 |     (Scan.option (Parse.name --| @{keyword ":"}) -- Parse.term >> uncurry functor_cmd);
 | 
| 40583 | 281 | |
| 40582 | 282 | end; |