| author | blanchet | 
| Fri, 04 Jan 2013 21:56:19 +0100 | |
| changeset 50732 | b2e7490a1b3d | 
| parent 47698 | 18202d3d5832 | 
| child 55414 | eab03e9cee8a | 
| permissions | -rw-r--r-- | 
| 37744 | 1 | (* Title: HOL/Tools/Quotient/quotient_term.ML | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 2 | Author: Cezary Kaliszyk and Christian Urban | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 3 | |
| 35788 | 4 | Constructs terms corresponding to goals from lifting theorems to | 
| 5 | quotient types. | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 6 | *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 7 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 8 | signature QUOTIENT_TERM = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 9 | sig | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 10 | exception LIFT_MATCH of string | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 11 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 12 | datatype flag = AbsF | RepF | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 13 | |
| 45797 | 14 | val absrep_fun: Proof.context -> flag -> typ * typ -> term | 
| 15 | val absrep_fun_chk: Proof.context -> flag -> typ * typ -> term | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 16 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 17 | (* Allows Nitpick to represent quotient types as single elements from raw type *) | 
| 45797 | 18 | val absrep_const_chk: Proof.context -> flag -> string -> term | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 19 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 20 | val equiv_relation: Proof.context -> typ * typ -> term | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 21 | val equiv_relation_chk: Proof.context -> typ * typ -> term | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 22 | |
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 23 | val get_rel_from_quot_thm: thm -> term | 
| 47504 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47308diff
changeset | 24 | val prove_quot_thm: Proof.context -> typ * typ -> thm | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 25 | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 26 | val regularize_trm: Proof.context -> term * term -> term | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 27 | val regularize_trm_chk: Proof.context -> term * term -> term | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 28 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 29 | val inj_repabs_trm: Proof.context -> term * term -> term | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 30 | val inj_repabs_trm_chk: Proof.context -> term * term -> term | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 31 | |
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 32 | val derive_qtyp: Proof.context -> typ list -> typ -> typ | 
| 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 33 | val derive_qtrm: Proof.context -> typ list -> term -> term | 
| 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 34 | val derive_rtyp: Proof.context -> typ list -> typ -> typ | 
| 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 35 | val derive_rtrm: Proof.context -> typ list -> term -> term | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 36 | end; | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 37 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 38 | structure Quotient_Term: QUOTIENT_TERM = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 39 | struct | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 40 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 41 | exception LIFT_MATCH of string | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 42 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 43 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 44 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 45 | (*** Aggregate Rep/Abs Function ***) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 46 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 47 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 48 | (* The flag RepF is for types in negative position; AbsF is for types | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 49 | in positive position. Because of this, function types need to be | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 50 | treated specially, since there the polarity changes. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 51 | *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 52 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 53 | datatype flag = AbsF | RepF | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 54 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 55 | fun negF AbsF = RepF | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 56 | | negF RepF = AbsF | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 57 | |
| 37677 | 58 | fun is_identity (Const (@{const_name id}, _)) = true
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 59 | | is_identity _ = false | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 60 | |
| 37677 | 61 | fun mk_identity ty = Const (@{const_name id}, ty --> ty)
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 62 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 63 | fun mk_fun_compose flag (trm1, trm2) = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 64 | case flag of | 
| 37677 | 65 |     AbsF => Const (@{const_name comp}, dummyT) $ trm1 $ trm2
 | 
| 66 |   | RepF => Const (@{const_name comp}, dummyT) $ trm2 $ trm1
 | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 67 | |
| 45796 | 68 | fun get_mapfun_data ctxt s = | 
| 69 | (case Symtab.lookup (Enriched_Type.entries ctxt) s of | |
| 45795 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 70 | SOME [map_data] => (case try dest_Const (#mapper map_data) of | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 71 | SOME (c, _) => (Const (c, dummyT), #variances map_data) | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 72 |     | NONE => raise LIFT_MATCH ("map function for type " ^ quote s ^ " is not a constant."))
 | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 73 |   | SOME _ => raise LIFT_MATCH ("map function for type " ^ quote s ^ " is non-singleton entry.")
 | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 74 |   | NONE => raise LIFT_MATCH ("No map function for type " ^ quote s ^ " found.")) 
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 75 | |
| 45796 | 76 | fun defined_mapfun_data ctxt s = | 
| 77 | Symtab.defined (Enriched_Type.entries ctxt) s | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 78 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 79 | (* looks up the (varified) rty and qty for | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 80 | a quotient definition | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 81 | *) | 
| 45796 | 82 | fun get_rty_qty ctxt s = | 
| 83 | let | |
| 84 | val thy = Proof_Context.theory_of ctxt | |
| 85 | in | |
| 86 | (case Quotient_Info.lookup_quotients_global thy s of | |
| 87 | SOME qdata => (#rtyp qdata, #qtyp qdata) | |
| 88 |     | NONE => raise LIFT_MATCH ("No quotient type " ^ quote s ^ " found."))
 | |
| 89 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 90 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 91 | (* matches a type pattern with a type *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 92 | fun match ctxt err ty_pat ty = | 
| 41444 | 93 | let | 
| 42361 | 94 | val thy = Proof_Context.theory_of ctxt | 
| 41444 | 95 | in | 
| 96 | Sign.typ_match thy (ty_pat, ty) Vartab.empty | |
| 97 | handle Type.TYPE_MATCH => err ctxt ty_pat ty | |
| 98 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 99 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 100 | (* produces the rep or abs constant for a qty *) | 
| 45797 | 101 | fun absrep_const ctxt flag qty_str = | 
| 41444 | 102 | let | 
| 45534 
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
 bulwahn parents: 
45340diff
changeset | 103 | (* FIXME *) | 
| 
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
 bulwahn parents: 
45340diff
changeset | 104 | fun mk_dummyT (Const (c, _)) = Const (c, dummyT) | 
| 45628 
f21eb7073895
in a local context, also the free variable case needs to be analysed default tip
 Christian Urban <urbanc@in.tum.de> parents: 
45534diff
changeset | 105 | | mk_dummyT (Free (c, _)) = Free (c, dummyT) | 
| 
f21eb7073895
in a local context, also the free variable case needs to be analysed default tip
 Christian Urban <urbanc@in.tum.de> parents: 
45534diff
changeset | 106 | | mk_dummyT _ = error "Expecting abs/rep term to be a constant or a free variable" | 
| 41444 | 107 | in | 
| 45534 
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
 bulwahn parents: 
45340diff
changeset | 108 | case Quotient_Info.lookup_abs_rep ctxt qty_str of | 
| 
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
 bulwahn parents: 
45340diff
changeset | 109 | SOME abs_rep => | 
| 
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
 bulwahn parents: 
45340diff
changeset | 110 | mk_dummyT (case flag of | 
| 
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
 bulwahn parents: 
45340diff
changeset | 111 | AbsF => #abs abs_rep | 
| 
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
 bulwahn parents: 
45340diff
changeset | 112 | | RepF => #rep abs_rep) | 
| 
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
 bulwahn parents: 
45340diff
changeset | 113 |       | NONE => error ("No abs/rep terms for " ^ quote qty_str)
 | 
| 41444 | 114 | end | 
| 45534 
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
 bulwahn parents: 
45340diff
changeset | 115 | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 116 | (* Lets Nitpick represent elements of quotient types as elements of the raw type *) | 
| 45797 | 117 | fun absrep_const_chk ctxt flag qty_str = | 
| 118 | Syntax.check_term ctxt (absrep_const ctxt flag qty_str) | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 119 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 120 | fun absrep_match_err ctxt ty_pat ty = | 
| 41444 | 121 | let | 
| 122 | val ty_pat_str = Syntax.string_of_typ ctxt ty_pat | |
| 123 | val ty_str = Syntax.string_of_typ ctxt ty | |
| 124 | in | |
| 125 | raise LIFT_MATCH (space_implode " " | |
| 126 | ["absrep_fun (Types ", quote ty_pat_str, "and", quote ty_str, " do not match.)"]) | |
| 127 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 128 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 129 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 130 | (** generation of an aggregate absrep function **) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 131 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 132 | (* - In case of equal types we just return the identity. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 133 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 134 | - In case of TFrees we also return the identity. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 135 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 136 | - In case of function types we recurse taking | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 137 | the polarity change into account. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 138 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 139 | - If the type constructors are equal, we recurse for the | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 140 | arguments and build the appropriate map function. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 141 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 142 | - If the type constructors are unequal, there must be an | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 143 | instance of quotient types: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 144 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 145 | - we first look up the corresponding rty_pat and qty_pat | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 146 | from the quotient definition; the arguments of qty_pat | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 147 | must be some distinct TVars | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 148 | - we then match the rty_pat with rty and qty_pat with qty; | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 149 | if matching fails the types do not correspond -> error | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 150 | - the matching produces two environments; we look up the | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 151 | assignments for the qty_pat variables and recurse on the | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 152 | assignments | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 153 | - we prefix the aggregate map function for the rty_pat, | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 154 | which is an abstraction over all type variables | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 155 | - finally we compose the result with the appropriate | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 156 | absrep function in case at least one argument produced | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 157 | a non-identity function / | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 158 | otherwise we just return the appropriate absrep | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 159 | function | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 160 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 161 | The composition is necessary for types like | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 162 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 163 |         ('a list) list / ('a foo) foo
 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 164 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 165 | The matching is necessary for types like | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 166 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 167 |         ('a * 'a) list / 'a bar
 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 168 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 169 | The test is necessary in order to eliminate superfluous | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 170 | identity maps. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 171 | *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 172 | |
| 45797 | 173 | fun absrep_fun ctxt flag (rty, qty) = | 
| 45340 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 174 | let | 
| 45795 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 175 | fun absrep_args tys tys' variances = | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 176 | let | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 177 | fun absrep_arg (types, (_, variant)) = | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 178 | (case variant of | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 179 | (false, false) => [] | 
| 45797 | 180 | | (true, false) => [(absrep_fun ctxt flag types)] | 
| 181 | | (false, true) => [(absrep_fun ctxt (negF flag) types)] | |
| 182 | | (true, true) => [(absrep_fun ctxt flag types),(absrep_fun ctxt (negF flag) types)]) | |
| 45795 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 183 | in | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 184 | maps absrep_arg ((tys ~~ tys') ~~ variances) | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 185 | end | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 186 | fun test_identities tys rtys' s s' = | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 187 | let | 
| 45797 | 188 | val args = map (absrep_fun ctxt flag) (tys ~~ rtys') | 
| 45795 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 189 | in | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 190 | if forall is_identity args | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 191 | then | 
| 45797 | 192 | absrep_const ctxt flag s' | 
| 45795 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 193 | else | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 194 |           raise LIFT_MATCH ("No map function for type " ^ quote s ^ " found.")
 | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 195 | end | 
| 45340 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 196 | in | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 197 | if rty = qty | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 198 | then mk_identity rty | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 199 | else | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 200 | case (rty, qty) of | 
| 45795 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 201 | (Type (s, tys), Type (s', tys')) => | 
| 45340 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 202 | if s = s' | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 203 | then | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 204 | let | 
| 45796 | 205 | val (map_fun, variances) = get_mapfun_data ctxt s | 
| 45795 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 206 | val args = absrep_args tys tys' variances | 
| 45340 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 207 | in | 
| 45795 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 208 | list_comb (map_fun, args) | 
| 45340 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 209 | end | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 210 | else | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 211 | let | 
| 45796 | 212 | val (Type (_, rtys), qty_pat) = get_rty_qty ctxt s' | 
| 45340 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 213 | val qtyenv = match ctxt absrep_match_err qty_pat qty | 
| 45795 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 214 | val rtys' = map (Envir.subst_type qtyenv) rtys | 
| 45340 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 215 | in | 
| 45796 | 216 | if not (defined_mapfun_data ctxt s) | 
| 45795 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 217 | then | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 218 | (* | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 219 | If we don't know a map function for the raw type, | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 220 | we are not necessarilly in troubles because | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 221 | it can still be the case we don't need the map | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 222 | function <=> all abs/rep functions are identities. | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 223 | *) | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 224 | test_identities tys rtys' s s' | 
| 45340 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 225 | else | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 226 | let | 
| 45796 | 227 | val (map_fun, variances) = get_mapfun_data ctxt s | 
| 45795 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 228 | val args = absrep_args tys rtys' variances | 
| 45340 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 229 | in | 
| 45795 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 230 | if forall is_identity args | 
| 45797 | 231 | then absrep_const ctxt flag s' | 
| 45795 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 232 | else | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 233 | let | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 234 | val result = list_comb (map_fun, args) | 
| 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 235 | in | 
| 45797 | 236 | mk_fun_compose flag (absrep_const ctxt flag s', result) | 
| 45795 
2d8949268303
maps are taken from enriched type infrastracture, rewritten lifting of constants, now we can lift even contravariant and co/contravariant types
 kuncar parents: 
45628diff
changeset | 237 | end | 
| 45340 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 238 | end | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 239 | end | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 240 | | (TFree x, TFree x') => | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 241 | if x = x' | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 242 | then mk_identity rty | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 243 | else raise (LIFT_MATCH "absrep_fun (frees)") | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 244 | | (TVar _, TVar _) => raise (LIFT_MATCH "absrep_fun (vars)") | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 245 | | _ => raise (LIFT_MATCH "absrep_fun (default)") | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 246 | end | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 247 | |
| 45797 | 248 | fun absrep_fun_chk ctxt flag (rty, qty) = | 
| 249 | absrep_fun ctxt flag (rty, qty) | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 250 | |> Syntax.check_term ctxt | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 251 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 252 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 253 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 254 | (*** Aggregate Equivalence Relation ***) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 255 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 256 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 257 | (* works very similar to the absrep generation, | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 258 | except there is no need for polarities | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 259 | *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 260 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 261 | (* instantiates TVars so that the term is of type ty *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 262 | fun force_typ ctxt trm ty = | 
| 41444 | 263 | let | 
| 42361 | 264 | val thy = Proof_Context.theory_of ctxt | 
| 41444 | 265 | val trm_ty = fastype_of trm | 
| 266 | val ty_inst = Sign.typ_match thy (trm_ty, ty) Vartab.empty | |
| 267 | in | |
| 268 | map_types (Envir.subst_type ty_inst) trm | |
| 269 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 270 | |
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 271 | fun is_eq (Const (@{const_name HOL.eq}, _)) = true
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 272 | | is_eq _ = false | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 273 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 274 | fun mk_rel_compose (trm1, trm2) = | 
| 35402 | 275 |   Const (@{const_abbrev "rel_conj"}, dummyT) $ trm1 $ trm2
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 276 | |
| 45340 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 277 | fun get_relmap thy s = | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 278 | (case Quotient_Info.lookup_quotmaps thy s of | 
| 45273 
728ed9d28c63
respecting isabelle's programming style in the quotient package by simplifying map_lookup function for data access
 bulwahn parents: 
45272diff
changeset | 279 | SOME map_data => Const (#relmap map_data, dummyT) | 
| 45279 | 280 |   | NONE => raise LIFT_MATCH ("get_relmap (no relation map function found for type " ^ s ^ ")"))
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 281 | |
| 45340 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 282 | fun get_equiv_rel thy s = | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 283 | (case Quotient_Info.lookup_quotients thy s of | 
| 45279 | 284 | SOME qdata => #equiv_rel qdata | 
| 47095 
b43ddeea727f
simplified code of generation of aggregate relations
 kuncar parents: 
46416diff
changeset | 285 |   | NONE => raise LIFT_MATCH ("get_equiv_rel (no quotient found for type " ^ s ^ ")"))
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 286 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 287 | fun equiv_match_err ctxt ty_pat ty = | 
| 41444 | 288 | let | 
| 289 | val ty_pat_str = Syntax.string_of_typ ctxt ty_pat | |
| 290 | val ty_str = Syntax.string_of_typ ctxt ty | |
| 291 | in | |
| 292 | raise LIFT_MATCH (space_implode " " | |
| 293 | ["equiv_relation (Types ", quote ty_pat_str, "and", quote ty_str, " do not match.)"]) | |
| 294 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 295 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 296 | (* builds the aggregate equivalence relation | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 297 | that will be the argument of Respects | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 298 | *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 299 | fun equiv_relation ctxt (rty, qty) = | 
| 45796 | 300 | if rty = qty | 
| 301 | then HOLogic.eq_const rty | |
| 302 | else | |
| 303 | case (rty, qty) of | |
| 304 | (Type (s, tys), Type (s', tys')) => | |
| 305 | if s = s' | |
| 306 | then | |
| 307 | let | |
| 308 | val args = map (equiv_relation ctxt) (tys ~~ tys') | |
| 309 | in | |
| 310 | list_comb (get_relmap ctxt s, args) | |
| 311 | end | |
| 312 | else | |
| 313 | let | |
| 47095 
b43ddeea727f
simplified code of generation of aggregate relations
 kuncar parents: 
46416diff
changeset | 314 | val (Type (_, rtys), qty_pat) = get_rty_qty ctxt s' | 
| 45796 | 315 | val qtyenv = match ctxt equiv_match_err qty_pat qty | 
| 47095 
b43ddeea727f
simplified code of generation of aggregate relations
 kuncar parents: 
46416diff
changeset | 316 | val rtys' = map (Envir.subst_type qtyenv) rtys | 
| 
b43ddeea727f
simplified code of generation of aggregate relations
 kuncar parents: 
46416diff
changeset | 317 | val args = map (equiv_relation ctxt) (tys ~~ rtys') | 
| 45796 | 318 | val eqv_rel = get_equiv_rel ctxt s' | 
| 319 |             val eqv_rel' = force_typ ctxt eqv_rel ([rty, rty] ---> @{typ bool})
 | |
| 320 | in | |
| 321 | if forall is_eq args | |
| 322 | then eqv_rel' | |
| 323 | else | |
| 324 | let | |
| 47095 
b43ddeea727f
simplified code of generation of aggregate relations
 kuncar parents: 
46416diff
changeset | 325 | val result = list_comb (get_relmap ctxt s, args) | 
| 45796 | 326 | in | 
| 327 | mk_rel_compose (result, eqv_rel') | |
| 328 | end | |
| 329 | end | |
| 330 | | _ => HOLogic.eq_const rty | |
| 331 | ||
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 332 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 333 | fun equiv_relation_chk ctxt (rty, qty) = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 334 | equiv_relation ctxt (rty, qty) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 335 | |> Syntax.check_term ctxt | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 336 | |
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 337 | (* generation of the Quotient theorem *) | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 338 | |
| 47106 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 339 | exception CODE_GEN of string | 
| 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 340 | |
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 341 | fun get_quot_thm ctxt s = | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 342 | let | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 343 | val thy = Proof_Context.theory_of ctxt | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 344 | in | 
| 47106 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 345 | (case Quotient_Info.lookup_quotients ctxt s of | 
| 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 346 | SOME qdata => Thm.transfer thy (#quot_thm qdata) | 
| 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 347 |     | NONE => raise CODE_GEN ("No quotient type " ^ quote s ^ " found."))
 | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 348 | end | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 349 | |
| 47106 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 350 | fun get_rel_quot_thm ctxt s = | 
| 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 351 | let | 
| 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 352 | val thy = Proof_Context.theory_of ctxt | 
| 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 353 | in | 
| 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 354 | (case Quotient_Info.lookup_quotmaps ctxt s of | 
| 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 355 | SOME map_data => Thm.transfer thy (#quot_thm map_data) | 
| 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 356 |     | NONE => raise CODE_GEN ("get_relmap (no relation map function found for type " ^ s ^ ")"))
 | 
| 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 357 | end | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 358 | |
| 47308 | 359 | fun is_id_quot thm = (prop_of thm = prop_of @{thm identity_quotient3})
 | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 360 | |
| 47698 | 361 | open Lifting_Util | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 362 | |
| 47698 | 363 | infix 0 MRSL | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 364 | |
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 365 | exception NOT_IMPL of string | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 366 | |
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 367 | fun get_rel_from_quot_thm quot_thm = | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 368 | let | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 369 | val (_ $ rel $ _ $ _) = (HOLogic.dest_Trueprop o prop_of) quot_thm | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 370 | in | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 371 | rel | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 372 | end | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 373 | |
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 374 | fun mk_quot_thm_compose (rel_quot_thm, quot_thm) = | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 375 | let | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 376 | val quot_thm_rel = get_rel_from_quot_thm quot_thm | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 377 | in | 
| 47308 | 378 |     if is_eq quot_thm_rel then [rel_quot_thm, quot_thm] MRSL @{thm OOO_eq_quotient3}
 | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 379 | else raise NOT_IMPL "nested quotients: not implemented yet" | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 380 | end | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 381 | |
| 47504 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47308diff
changeset | 382 | fun prove_quot_thm ctxt (rty, qty) = | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 383 | if rty = qty | 
| 47308 | 384 |   then @{thm identity_quotient3}
 | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 385 | else | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 386 | case (rty, qty) of | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 387 | (Type (s, tys), Type (s', tys')) => | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 388 | if s = s' | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 389 | then | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 390 | let | 
| 47504 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47308diff
changeset | 391 | val args = map (prove_quot_thm ctxt) (tys ~~ tys') | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 392 | in | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 393 | args MRSL (get_rel_quot_thm ctxt s) | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 394 | end | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 395 | else | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 396 | let | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 397 | val (Type (_, rtys), qty_pat) = get_rty_qty ctxt s' | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 398 | val qtyenv = match ctxt equiv_match_err qty_pat qty | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 399 | val rtys' = map (Envir.subst_type qtyenv) rtys | 
| 47504 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47308diff
changeset | 400 | val args = map (prove_quot_thm ctxt) (tys ~~ rtys') | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 401 | val quot_thm = get_quot_thm ctxt s' | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 402 | in | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 403 | if forall is_id_quot args | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 404 | then | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 405 | quot_thm | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 406 | else | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 407 | let | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 408 | val rel_quot_thm = args MRSL (get_rel_quot_thm ctxt s) | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 409 | in | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 410 | mk_quot_thm_compose (rel_quot_thm, quot_thm) | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 411 | end | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 412 | end | 
| 47308 | 413 |     | _ => @{thm identity_quotient3}
 | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47095diff
changeset | 414 | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 415 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 416 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 417 | (*** Regularization ***) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 418 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 419 | (* Regularizing an rtrm means: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 420 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 421 | - Quantifiers over types that need lifting are replaced | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 422 | by bounded quantifiers, for example: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 423 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 424 | All P ----> All (Respects R) P | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 425 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 426 | where the aggregate relation R is given by the rty and qty; | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 427 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 428 | - Abstractions over types that need lifting are replaced | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 429 | by bounded abstractions, for example: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 430 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 431 | %x. P ----> Ball (Respects R) %x. P | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 432 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 433 | - Equalities over types that need lifting are replaced by | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 434 | corresponding equivalence relations, for example: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 435 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 436 | A = B ----> R A B | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 437 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 438 | or | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 439 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 440 | A = B ----> (R ===> R) A B | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 441 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 442 | for more complicated types of A and B | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 443 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 444 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 445 | The regularize_trm accepts raw theorems in which equalities | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 446 | and quantifiers match exactly the ones in the lifted theorem | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 447 | but also accepts partially regularized terms. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 448 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 449 | This means that the raw theorems can have: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 450 | Ball (Respects R), Bex (Respects R), Bex1_rel (Respects R), Babs, R | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 451 | in the places where: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 452 | All, Ex, Ex1, %, (op =) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 453 | is required the lifted theorem. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 454 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 455 | *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 456 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 457 | val mk_babs = Const (@{const_name Babs}, dummyT)
 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 458 | val mk_ball = Const (@{const_name Ball}, dummyT)
 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 459 | val mk_bex  = Const (@{const_name Bex}, dummyT)
 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 460 | val mk_bex1_rel = Const (@{const_name Bex1_rel}, dummyT)
 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 461 | val mk_resp = Const (@{const_name Respects}, dummyT)
 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 462 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 463 | (* - applies f to the subterm of an abstraction, | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 464 | otherwise to the given term, | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 465 | - used by regularize, therefore abstracted | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 466 | variables do not have to be treated specially | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 467 | *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 468 | fun apply_subt f (trm1, trm2) = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 469 | case (trm1, trm2) of | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 470 | (Abs (x, T, t), Abs (_ , _, t')) => Abs (x, T, f (t, t')) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 471 | | _ => f (trm1, trm2) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 472 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 473 | fun term_mismatch str ctxt t1 t2 = | 
| 41444 | 474 | let | 
| 475 | val t1_str = Syntax.string_of_term ctxt t1 | |
| 476 | val t2_str = Syntax.string_of_term ctxt t2 | |
| 477 | val t1_ty_str = Syntax.string_of_typ ctxt (fastype_of t1) | |
| 478 | val t2_ty_str = Syntax.string_of_typ ctxt (fastype_of t2) | |
| 479 | in | |
| 480 | raise LIFT_MATCH (cat_lines [str, t1_str ^ "::" ^ t1_ty_str, t2_str ^ "::" ^ t2_ty_str]) | |
| 481 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 482 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 483 | (* the major type of All and Ex quantifiers *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 484 | fun qnt_typ ty = domain_type (domain_type ty) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 485 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 486 | (* Checks that two types match, for example: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 487 | rty -> rty matches qty -> qty *) | 
| 45280 | 488 | fun matches_typ ctxt rT qT = | 
| 45340 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 489 | let | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 490 | val thy = Proof_Context.theory_of ctxt | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 491 | in | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 492 | if rT = qT then true | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 493 | else | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 494 | (case (rT, qT) of | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 495 | (Type (rs, rtys), Type (qs, qtys)) => | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 496 | if rs = qs then | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 497 | if length rtys <> length qtys then false | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 498 | else forall (fn x => x = true) (map2 (matches_typ ctxt) rtys qtys) | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 499 | else | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 500 | (case Quotient_Info.lookup_quotients_global thy qs of | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 501 | SOME quotinfo => Sign.typ_instance thy (rT, #rtyp quotinfo) | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 502 | | NONE => false) | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 503 | | _ => false) | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 504 | end | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 505 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 506 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 507 | (* produces a regularized version of rtrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 508 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 509 | - the result might contain dummyTs | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 510 | |
| 38718 | 511 | - for regularization we do not need any | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 512 | special treatment of bound variables | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 513 | *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 514 | fun regularize_trm ctxt (rtrm, qtrm) = | 
| 45280 | 515 | (case (rtrm, qtrm) of | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 516 | (Abs (x, ty, t), Abs (_, ty', t')) => | 
| 41444 | 517 | let | 
| 518 | val subtrm = Abs(x, ty, regularize_trm ctxt (t, t')) | |
| 519 | in | |
| 520 | if ty = ty' then subtrm | |
| 521 | else mk_babs $ (mk_resp $ equiv_relation ctxt (ty, ty')) $ subtrm | |
| 522 | end | |
| 45280 | 523 | |
| 37677 | 524 |   | (Const (@{const_name Babs}, T) $ resrel $ (t as (Abs (_, ty, _))), t' as (Abs (_, ty', _))) =>
 | 
| 41444 | 525 | let | 
| 526 | val subtrm = regularize_trm ctxt (t, t') | |
| 527 | val needres = mk_resp $ equiv_relation_chk ctxt (ty, ty') | |
| 528 | in | |
| 529 | if resrel <> needres | |
| 530 | then term_mismatch "regularize (Babs)" ctxt resrel needres | |
| 531 | else mk_babs $ resrel $ subtrm | |
| 532 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 533 | |
| 37677 | 534 |   | (Const (@{const_name All}, ty) $ t, Const (@{const_name All}, ty') $ t') =>
 | 
| 41444 | 535 | let | 
| 536 | val subtrm = apply_subt (regularize_trm ctxt) (t, t') | |
| 537 | in | |
| 538 |         if ty = ty' then Const (@{const_name All}, ty) $ subtrm
 | |
| 539 | else mk_ball $ (mk_resp $ equiv_relation ctxt (qnt_typ ty, qnt_typ ty')) $ subtrm | |
| 540 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 541 | |
| 37677 | 542 |   | (Const (@{const_name Ex}, ty) $ t, Const (@{const_name Ex}, ty') $ t') =>
 | 
| 41444 | 543 | let | 
| 544 | val subtrm = apply_subt (regularize_trm ctxt) (t, t') | |
| 545 | in | |
| 546 |         if ty = ty' then Const (@{const_name Ex}, ty) $ subtrm
 | |
| 547 | else mk_bex $ (mk_resp $ equiv_relation ctxt (qnt_typ ty, qnt_typ ty')) $ subtrm | |
| 548 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 549 | |
| 37677 | 550 |   | (Const (@{const_name Ex1}, ty) $ (Abs (_, _,
 | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38718diff
changeset | 551 |       (Const (@{const_name HOL.conj}, _) $ (Const (@{const_name Set.member}, _) $ _ $
 | 
| 37677 | 552 |         (Const (@{const_name Respects}, _) $ resrel)) $ (t $ _)))),
 | 
| 553 |      Const (@{const_name Ex1}, ty') $ t') =>
 | |
| 41444 | 554 | let | 
| 555 | val t_ = incr_boundvars (~1) t | |
| 556 | val subtrm = apply_subt (regularize_trm ctxt) (t_, t') | |
| 557 | val needrel = equiv_relation_chk ctxt (qnt_typ ty, qnt_typ ty') | |
| 558 | in | |
| 559 | if resrel <> needrel | |
| 560 | then term_mismatch "regularize (Bex1)" ctxt resrel needrel | |
| 561 | else mk_bex1_rel $ resrel $ subtrm | |
| 562 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 563 | |
| 38558 | 564 |   | (Const (@{const_name Ex1}, ty) $ t, Const (@{const_name Ex1}, ty') $ t') =>
 | 
| 41444 | 565 | let | 
| 566 | val subtrm = apply_subt (regularize_trm ctxt) (t, t') | |
| 567 | in | |
| 568 |         if ty = ty' then Const (@{const_name Ex1}, ty) $ subtrm
 | |
| 569 | else mk_bex1_rel $ (equiv_relation ctxt (qnt_typ ty, qnt_typ ty')) $ subtrm | |
| 570 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 571 | |
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 572 |   | (Const (@{const_name Ball}, ty) $ (Const (@{const_name Respects}, _) $ resrel) $ t,
 | 
| 38558 | 573 |      Const (@{const_name All}, ty') $ t') =>
 | 
| 41444 | 574 | let | 
| 575 | val subtrm = apply_subt (regularize_trm ctxt) (t, t') | |
| 576 | val needrel = equiv_relation_chk ctxt (qnt_typ ty, qnt_typ ty') | |
| 577 | in | |
| 578 | if resrel <> needrel | |
| 579 | then term_mismatch "regularize (Ball)" ctxt resrel needrel | |
| 580 | else mk_ball $ (mk_resp $ resrel) $ subtrm | |
| 581 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 582 | |
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 583 |   | (Const (@{const_name Bex}, ty) $ (Const (@{const_name Respects}, _) $ resrel) $ t,
 | 
| 38558 | 584 |      Const (@{const_name Ex}, ty') $ t') =>
 | 
| 41444 | 585 | let | 
| 586 | val subtrm = apply_subt (regularize_trm ctxt) (t, t') | |
| 587 | val needrel = equiv_relation_chk ctxt (qnt_typ ty, qnt_typ ty') | |
| 588 | in | |
| 589 | if resrel <> needrel | |
| 590 | then term_mismatch "regularize (Bex)" ctxt resrel needrel | |
| 591 | else mk_bex $ (mk_resp $ resrel) $ subtrm | |
| 592 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 593 | |
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 594 |   | (Const (@{const_name Bex1_rel}, ty) $ resrel $ t, Const (@{const_name Ex1}, ty') $ t') =>
 | 
| 41444 | 595 | let | 
| 596 | val subtrm = apply_subt (regularize_trm ctxt) (t, t') | |
| 597 | val needrel = equiv_relation_chk ctxt (qnt_typ ty, qnt_typ ty') | |
| 598 | in | |
| 599 | if resrel <> needrel | |
| 600 | then term_mismatch "regularize (Bex1_res)" ctxt resrel needrel | |
| 601 | else mk_bex1_rel $ resrel $ subtrm | |
| 602 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 603 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 604 | | (* equalities need to be replaced by appropriate equivalence relations *) | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 605 |     (Const (@{const_name HOL.eq}, ty), Const (@{const_name HOL.eq}, ty')) =>
 | 
| 41444 | 606 | if ty = ty' then rtrm | 
| 607 | else equiv_relation ctxt (domain_type ty, domain_type ty') | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 608 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 609 | | (* in this case we just check whether the given equivalence relation is correct *) | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 610 |     (rel, Const (@{const_name HOL.eq}, ty')) =>
 | 
| 41444 | 611 | let | 
| 612 | val rel_ty = fastype_of rel | |
| 613 | val rel' = equiv_relation_chk ctxt (domain_type rel_ty, domain_type ty') | |
| 614 | in | |
| 615 | if rel' aconv rel then rtrm | |
| 616 | else term_mismatch "regularize (relation mismatch)" ctxt rel rel' | |
| 617 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 618 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 619 | | (_, Const _) => | 
| 41444 | 620 | let | 
| 42361 | 621 | val thy = Proof_Context.theory_of ctxt | 
| 45280 | 622 | fun same_const (Const (s, T)) (Const (s', T')) = s = s' andalso matches_typ ctxt T T' | 
| 41444 | 623 | | same_const _ _ = false | 
| 624 | in | |
| 625 | if same_const rtrm qtrm then rtrm | |
| 626 | else | |
| 627 | let | |
| 45279 | 628 | val rtrm' = | 
| 45340 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 629 | (case Quotient_Info.lookup_quotconsts_global thy qtrm of | 
| 45279 | 630 | SOME qconst_info => #rconst qconst_info | 
| 631 | | NONE => term_mismatch "regularize (constant not found)" ctxt rtrm qtrm) | |
| 41444 | 632 | in | 
| 633 | if Pattern.matches thy (rtrm', rtrm) | |
| 634 | then rtrm else term_mismatch "regularize (constant mismatch)" ctxt rtrm qtrm | |
| 635 | end | |
| 636 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 637 | |
| 37591 | 638 |   | (((t1 as Const (@{const_name prod_case}, _)) $ Abs (v1, ty, Abs(v1', ty', s1))),
 | 
| 639 |      ((t2 as Const (@{const_name prod_case}, _)) $ Abs (v2, _ , Abs(v2', _  , s2)))) =>
 | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 640 | regularize_trm ctxt (t1, t2) $ Abs (v1, ty, Abs (v1', ty', regularize_trm ctxt (s1, s2))) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 641 | |
| 37591 | 642 |   | (((t1 as Const (@{const_name prod_case}, _)) $ Abs (v1, ty, s1)),
 | 
| 643 |      ((t2 as Const (@{const_name prod_case}, _)) $ Abs (v2, _ , s2))) =>
 | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 644 | regularize_trm ctxt (t1, t2) $ Abs (v1, ty, regularize_trm ctxt (s1, s2)) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 645 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 646 | | (t1 $ t2, t1' $ t2') => | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 647 | regularize_trm ctxt (t1, t1') $ regularize_trm ctxt (t2, t2') | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 648 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 649 | | (Bound i, Bound i') => | 
| 41444 | 650 | if i = i' then rtrm | 
| 651 | else raise (LIFT_MATCH "regularize (bounds mismatch)") | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 652 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 653 | | _ => | 
| 41444 | 654 | let | 
| 655 | val rtrm_str = Syntax.string_of_term ctxt rtrm | |
| 656 | val qtrm_str = Syntax.string_of_term ctxt qtrm | |
| 657 | in | |
| 658 |         raise (LIFT_MATCH ("regularize failed (default: " ^ rtrm_str ^ "," ^ qtrm_str ^ ")"))
 | |
| 45280 | 659 | end) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 660 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 661 | fun regularize_trm_chk ctxt (rtrm, qtrm) = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 662 | regularize_trm ctxt (rtrm, qtrm) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 663 | |> Syntax.check_term ctxt | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 664 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 665 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 666 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 667 | (*** Rep/Abs Injection ***) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 668 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 669 | (* | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 670 | Injection of Rep/Abs means: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 671 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 672 | For abstractions: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 673 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 674 | * If the type of the abstraction needs lifting, then we add Rep/Abs | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 675 | around the abstraction; otherwise we leave it unchanged. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 676 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 677 | For applications: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 678 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 679 | * If the application involves a bounded quantifier, we recurse on | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 680 | the second argument. If the application is a bounded abstraction, | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 681 | we always put an Rep/Abs around it (since bounded abstractions | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 682 | are assumed to always need lifting). Otherwise we recurse on both | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 683 | arguments. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 684 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 685 | For constants: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 686 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 687 | * If the constant is (op =), we leave it always unchanged. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 688 | Otherwise the type of the constant needs lifting, we put | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 689 | and Rep/Abs around it. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 690 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 691 | For free variables: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 692 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 693 | * We put a Rep/Abs around it if the type needs lifting. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 694 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 695 | Vars case cannot occur. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 696 | *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 697 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 698 | fun mk_repabs ctxt (T, T') trm = | 
| 45797 | 699 | absrep_fun ctxt RepF (T, T') $ (absrep_fun ctxt AbsF (T, T') $ trm) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 700 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 701 | fun inj_repabs_err ctxt msg rtrm qtrm = | 
| 41444 | 702 | let | 
| 703 | val rtrm_str = Syntax.string_of_term ctxt rtrm | |
| 704 | val qtrm_str = Syntax.string_of_term ctxt qtrm | |
| 705 | in | |
| 706 | raise LIFT_MATCH (space_implode " " [msg, quote rtrm_str, "and", quote qtrm_str]) | |
| 707 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 708 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 709 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 710 | (* bound variables need to be treated properly, | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 711 | as the type of subterms needs to be calculated *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 712 | fun inj_repabs_trm ctxt (rtrm, qtrm) = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 713 | case (rtrm, qtrm) of | 
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 714 |     (Const (@{const_name Ball}, T) $ r $ t, Const (@{const_name All}, _) $ t') =>
 | 
| 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 715 |        Const (@{const_name Ball}, T) $ r $ (inj_repabs_trm ctxt (t, t'))
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 716 | |
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 717 |   | (Const (@{const_name Bex}, T) $ r $ t, Const (@{const_name Ex}, _) $ t') =>
 | 
| 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 718 |        Const (@{const_name Bex}, T) $ r $ (inj_repabs_trm ctxt (t, t'))
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 719 | |
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 720 |   | (Const (@{const_name Babs}, T) $ r $ t, t' as (Abs _)) =>
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 721 | let | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 722 | val rty = fastype_of rtrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 723 | val qty = fastype_of qtrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 724 | in | 
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 725 |         mk_repabs ctxt (rty, qty) (Const (@{const_name Babs}, T) $ r $ (inj_repabs_trm ctxt (t, t')))
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 726 | end | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 727 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 728 | | (Abs (x, T, t), Abs (x', T', t')) => | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 729 | let | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 730 | val rty = fastype_of rtrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 731 | val qty = fastype_of qtrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 732 | val (y, s) = Term.dest_abs (x, T, t) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 733 | val (_, s') = Term.dest_abs (x', T', t') | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 734 | val yvar = Free (y, T) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 735 | val result = Term.lambda_name (y, yvar) (inj_repabs_trm ctxt (s, s')) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 736 | in | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 737 | if rty = qty then result | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 738 | else mk_repabs ctxt (rty, qty) result | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 739 | end | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 740 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 741 | | (t $ s, t' $ s') => | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 742 | (inj_repabs_trm ctxt (t, t')) $ (inj_repabs_trm ctxt (s, s')) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 743 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 744 | | (Free (_, T), Free (_, T')) => | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 745 | if T = T' then rtrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 746 | else mk_repabs ctxt (T, T') rtrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 747 | |
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 748 |   | (_, Const (@{const_name HOL.eq}, _)) => rtrm
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 749 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 750 | | (_, Const (_, T')) => | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 751 | let | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 752 | val rty = fastype_of rtrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 753 | in | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 754 | if rty = T' then rtrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 755 | else mk_repabs ctxt (rty, T') rtrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 756 | end | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 757 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 758 | | _ => inj_repabs_err ctxt "injection (default):" rtrm qtrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 759 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 760 | fun inj_repabs_trm_chk ctxt (rtrm, qtrm) = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 761 | inj_repabs_trm ctxt (rtrm, qtrm) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 762 | |> Syntax.check_term ctxt | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 763 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 764 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 765 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 766 | (*** Wrapper for automatically transforming an rthm into a qthm ***) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 767 | |
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 768 | (* substitutions functions for r/q-types and | 
| 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 769 | r/q-constants, respectively | 
| 37560 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 770 | *) | 
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 771 | fun subst_typ ctxt ty_subst rty = | 
| 37560 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 772 | case rty of | 
| 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 773 | Type (s, rtys) => | 
| 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 774 | let | 
| 42361 | 775 | val thy = Proof_Context.theory_of ctxt | 
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 776 | val rty' = Type (s, map (subst_typ ctxt ty_subst) rtys) | 
| 37560 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 777 | |
| 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 778 | fun matches [] = rty' | 
| 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 779 | | matches ((rty, qty)::tail) = | 
| 45280 | 780 | (case try (Sign.typ_match thy (rty, rty')) Vartab.empty of | 
| 37560 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 781 | NONE => matches tail | 
| 46416 
5f5665a0b973
Make automatic derivation of raw/quotient types more greedy to allow descending and quot_lifted for compound quotients.
 Cezary Kaliszyk <cezarykaliszyk@gmail.com> parents: 
45797diff
changeset | 782 | | SOME inst => subst_typ ctxt ty_subst (Envir.subst_type inst qty)) | 
| 37560 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 783 | in | 
| 41444 | 784 | matches ty_subst | 
| 785 | end | |
| 37560 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 786 | | _ => rty | 
| 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 787 | |
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 788 | fun subst_trm ctxt ty_subst trm_subst rtrm = | 
| 37560 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 789 | case rtrm of | 
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 790 | t1 $ t2 => (subst_trm ctxt ty_subst trm_subst t1) $ (subst_trm ctxt ty_subst trm_subst t2) | 
| 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 791 | | Abs (x, ty, t) => Abs (x, subst_typ ctxt ty_subst ty, subst_trm ctxt ty_subst trm_subst t) | 
| 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 792 | | Free(n, ty) => Free(n, subst_typ ctxt ty_subst ty) | 
| 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 793 | | Var(n, ty) => Var(n, subst_typ ctxt ty_subst ty) | 
| 37560 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 794 | | Bound i => Bound i | 
| 41444 | 795 | | Const (a, ty) => | 
| 37560 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 796 | let | 
| 42361 | 797 | val thy = Proof_Context.theory_of ctxt | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 798 | |
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 799 | fun matches [] = Const (a, subst_typ ctxt ty_subst ty) | 
| 37560 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 800 | | matches ((rconst, qconst)::tail) = | 
| 45280 | 801 | (case try (Pattern.match thy (rconst, rtrm)) (Vartab.empty, Vartab.empty) of | 
| 37560 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 802 | NONE => matches tail | 
| 46416 
5f5665a0b973
Make automatic derivation of raw/quotient types more greedy to allow descending and quot_lifted for compound quotients.
 Cezary Kaliszyk <cezarykaliszyk@gmail.com> parents: 
45797diff
changeset | 803 | | SOME inst => subst_trm ctxt ty_subst trm_subst (Envir.subst_term inst qconst)) | 
| 37560 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 804 | in | 
| 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 805 | matches trm_subst | 
| 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 806 | end | 
| 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 807 | |
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 808 | (* generate type and term substitutions out of the | 
| 41444 | 809 | qtypes involved in a quotient; the direction flag | 
| 810 | indicates in which direction the substitutions work: | |
| 811 | ||
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 812 | true: quotient -> raw | 
| 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 813 | false: raw -> quotient | 
| 37560 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 814 | *) | 
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 815 | fun mk_ty_subst qtys direction ctxt = | 
| 41444 | 816 | let | 
| 42361 | 817 | val thy = Proof_Context.theory_of ctxt | 
| 41444 | 818 | in | 
| 45279 | 819 | Quotient_Info.dest_quotients ctxt | 
| 41444 | 820 | |> map (fn x => (#rtyp x, #qtyp x)) | 
| 821 | |> filter (fn (_, qty) => member (Sign.typ_instance thy o swap) qtys qty) | |
| 822 | |> map (if direction then swap else I) | |
| 823 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 824 | |
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 825 | fun mk_trm_subst qtys direction ctxt = | 
| 41444 | 826 | let | 
| 827 | val subst_typ' = subst_typ ctxt (mk_ty_subst qtys direction ctxt) | |
| 828 | fun proper (t1, t2) = subst_typ' (fastype_of t1) = fastype_of t2 | |
| 37563 
6cf28a1dfd75
Add reverse lifting flag to automated theorem derivation
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
37560diff
changeset | 829 | |
| 41444 | 830 | val const_substs = | 
| 45279 | 831 | Quotient_Info.dest_quotconsts ctxt | 
| 41444 | 832 | |> map (fn x => (#rconst x, #qconst x)) | 
| 833 | |> map (if direction then swap else I) | |
| 37560 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 834 | |
| 41444 | 835 | val rel_substs = | 
| 45279 | 836 | Quotient_Info.dest_quotients ctxt | 
| 41444 | 837 | |> map (fn x => (#equiv_rel x, HOLogic.eq_const (#qtyp x))) | 
| 838 | |> map (if direction then swap else I) | |
| 839 | in | |
| 840 | filter proper (const_substs @ rel_substs) | |
| 841 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 842 | |
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 843 | |
| 37560 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 844 | (* derives a qtyp and qtrm out of a rtyp and rtrm, | 
| 41444 | 845 | respectively | 
| 37560 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 846 | *) | 
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 847 | fun derive_qtyp ctxt qtys rty = | 
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 848 | subst_typ ctxt (mk_ty_subst qtys false ctxt) rty | 
| 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 849 | |
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 850 | fun derive_qtrm ctxt qtys rtrm = | 
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 851 | subst_trm ctxt (mk_ty_subst qtys false ctxt) (mk_trm_subst qtys false ctxt) rtrm | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 852 | |
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 853 | (* derives a rtyp and rtrm out of a qtyp and qtrm, | 
| 41444 | 854 | respectively | 
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 855 | *) | 
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 856 | fun derive_rtyp ctxt qtys qty = | 
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 857 | subst_typ ctxt (mk_ty_subst qtys true ctxt) qty | 
| 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 858 | |
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 859 | fun derive_rtrm ctxt qtys qtrm = | 
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 860 | subst_trm ctxt (mk_ty_subst qtys true ctxt) (mk_trm_subst qtys true ctxt) qtrm | 
| 37560 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 861 | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 862 | |
| 45279 | 863 | end; |