| author | wenzelm | 
| Mon, 23 Mar 2015 17:01:47 +0100 | |
| changeset 59785 | 4e6ab5831cc0 | 
| parent 59582 | 0fbed69ff081 | 
| child 59848 | 18c21d5c9138 | 
| 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 = | 
| 55467 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 blanchet parents: 
55414diff
changeset | 69 | (case Symtab.lookup (Functor.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 | |
| 55467 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 blanchet parents: 
55414diff
changeset | 76 | fun defined_mapfun_data ctxt = Symtab.defined (Functor.entries ctxt) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 77 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 78 | (* looks up the (varified) rty and qty for | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 79 | a quotient definition | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 80 | *) | 
| 45796 | 81 | fun get_rty_qty ctxt s = | 
| 82 | let | |
| 83 | val thy = Proof_Context.theory_of ctxt | |
| 84 | in | |
| 85 | (case Quotient_Info.lookup_quotients_global thy s of | |
| 86 | SOME qdata => (#rtyp qdata, #qtyp qdata) | |
| 87 |     | NONE => raise LIFT_MATCH ("No quotient type " ^ quote s ^ " found."))
 | |
| 88 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 89 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 90 | (* matches a type pattern with a type *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 91 | fun match ctxt err ty_pat ty = | 
| 41444 | 92 | let | 
| 42361 | 93 | val thy = Proof_Context.theory_of ctxt | 
| 41444 | 94 | in | 
| 95 | Sign.typ_match thy (ty_pat, ty) Vartab.empty | |
| 96 | handle Type.TYPE_MATCH => err ctxt ty_pat ty | |
| 97 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 98 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 99 | (* produces the rep or abs constant for a qty *) | 
| 45797 | 100 | fun absrep_const ctxt flag qty_str = | 
| 41444 | 101 | 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 | 102 | (* 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 | 103 | 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 | 104 | | 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 | 105 | | mk_dummyT _ = error "Expecting abs/rep term to be a constant or a free variable" | 
| 41444 | 106 | 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 | 107 | 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 | 108 | 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 | 109 | 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 | 110 | 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 | 111 | | 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 | 112 |       | NONE => error ("No abs/rep terms for " ^ quote qty_str)
 | 
| 41444 | 113 | 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 | 114 | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 115 | (* Lets Nitpick represent elements of quotient types as elements of the raw type *) | 
| 45797 | 116 | fun absrep_const_chk ctxt flag qty_str = | 
| 117 | 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 | 118 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 119 | fun absrep_match_err ctxt ty_pat ty = | 
| 41444 | 120 | let | 
| 121 | val ty_pat_str = Syntax.string_of_typ ctxt ty_pat | |
| 122 | val ty_str = Syntax.string_of_typ ctxt ty | |
| 123 | in | |
| 124 | raise LIFT_MATCH (space_implode " " | |
| 125 | ["absrep_fun (Types ", quote ty_pat_str, "and", quote ty_str, " do not match.)"]) | |
| 126 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 127 | |
| 
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 | (** generation of an aggregate absrep function **) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 130 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 131 | (* - 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 | 132 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 133 | - 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 | 134 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 135 | - In case of function types we recurse taking | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 136 | the polarity change into account. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 137 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 138 | - 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 | 139 | arguments and build the appropriate map function. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 140 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 141 | - 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 | 142 | instance of quotient types: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 143 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 144 | - 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 | 145 | 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 | 146 | must be some distinct TVars | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 147 | - 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 | 148 | 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 | 149 | - 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 | 150 | 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 | 151 | assignments | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 152 | - 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 | 153 | which is an abstraction over all type variables | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 154 | - finally we compose the result with the appropriate | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 155 | 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 | 156 | a non-identity function / | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 157 | otherwise we just return the appropriate absrep | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 158 | function | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 159 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 160 | The composition is necessary for types like | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 161 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 162 |         ('a list) list / ('a foo) foo
 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 163 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 164 | The matching is necessary for types like | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 165 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 166 |         ('a * 'a) list / 'a bar
 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 167 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 168 | 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 | 169 | identity maps. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 170 | *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 171 | |
| 45797 | 172 | 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 | 173 | 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 | 174 | 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 | 175 | 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 | 176 | 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 | 177 | (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 | 178 | (false, false) => [] | 
| 45797 | 179 | | (true, false) => [(absrep_fun ctxt flag types)] | 
| 180 | | (false, true) => [(absrep_fun ctxt (negF flag) types)] | |
| 181 | | (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 | 182 | 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 | 183 | 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 | 184 | 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 | 185 | 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 | 186 | let | 
| 45797 | 187 | 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 | 188 | 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 | 189 | 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 | 190 | then | 
| 45797 | 191 | 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 | 192 | 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 | 193 |           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 | 194 | end | 
| 45340 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 195 | in | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 196 | 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 | 197 | 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 | 198 | else | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 199 | 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 | 200 | (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 | 201 | 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 | 202 | then | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 203 | let | 
| 45796 | 204 | 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 | 205 | 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 | 206 | 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 | 207 | 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 | 208 | end | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 209 | else | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 210 | let | 
| 45796 | 211 | 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 | 212 | 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 | 213 | 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 | 214 | in | 
| 45796 | 215 | 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 | 216 | 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 | 217 | (* | 
| 
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 | 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 | 219 | 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 | 220 | 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 | 221 | 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 | 222 | *) | 
| 
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 | 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 | 224 | else | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 225 | let | 
| 45796 | 226 | 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 | 227 | 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 | 228 | 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 | 229 | if forall is_identity args | 
| 45797 | 230 | 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 | 231 | 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 | 232 | 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 | 233 | 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 | 234 | in | 
| 45797 | 235 | 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 | 236 | end | 
| 45340 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 237 | end | 
| 
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 | | (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 | 240 | 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 | 241 | 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 | 242 | 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 | 243 | | (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 | 244 | | _ => 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 | 245 | end | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 246 | |
| 45797 | 247 | fun absrep_fun_chk ctxt flag (rty, qty) = | 
| 248 | absrep_fun ctxt flag (rty, qty) | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 249 | |> Syntax.check_term ctxt | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 250 | |
| 
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 | (*** Aggregate Equivalence Relation ***) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 254 | |
| 
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 | (* works very similar to the absrep generation, | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 257 | except there is no need for polarities | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 258 | *) | 
| 
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 | (* 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 | 261 | fun force_typ ctxt trm ty = | 
| 41444 | 262 | let | 
| 42361 | 263 | val thy = Proof_Context.theory_of ctxt | 
| 41444 | 264 | val trm_ty = fastype_of trm | 
| 265 | val ty_inst = Sign.typ_match thy (trm_ty, ty) Vartab.empty | |
| 266 | in | |
| 267 | map_types (Envir.subst_type ty_inst) trm | |
| 268 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 269 | |
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 270 | 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 | 271 | | is_eq _ = false | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 272 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 273 | fun mk_rel_compose (trm1, trm2) = | 
| 35402 | 274 |   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 | 275 | |
| 45340 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 276 | 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 | 277 | (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 | 278 | SOME map_data => Const (#relmap map_data, dummyT) | 
| 45279 | 279 |   | 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 | 280 | |
| 45340 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 281 | 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 | 282 | (case Quotient_Info.lookup_quotients thy s of | 
| 45279 | 283 | SOME qdata => #equiv_rel qdata | 
| 47095 
b43ddeea727f
simplified code of generation of aggregate relations
 kuncar parents: 
46416diff
changeset | 284 |   | 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 | 285 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 286 | fun equiv_match_err ctxt ty_pat ty = | 
| 41444 | 287 | let | 
| 288 | val ty_pat_str = Syntax.string_of_typ ctxt ty_pat | |
| 289 | val ty_str = Syntax.string_of_typ ctxt ty | |
| 290 | in | |
| 291 | raise LIFT_MATCH (space_implode " " | |
| 292 | ["equiv_relation (Types ", quote ty_pat_str, "and", quote ty_str, " do not match.)"]) | |
| 293 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 294 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 295 | (* builds the aggregate equivalence relation | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 296 | that will be the argument of Respects | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 297 | *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 298 | fun equiv_relation ctxt (rty, qty) = | 
| 45796 | 299 | if rty = qty | 
| 300 | then HOLogic.eq_const rty | |
| 301 | else | |
| 302 | case (rty, qty) of | |
| 303 | (Type (s, tys), Type (s', tys')) => | |
| 304 | if s = s' | |
| 305 | then | |
| 306 | let | |
| 307 | val args = map (equiv_relation ctxt) (tys ~~ tys') | |
| 308 | in | |
| 309 | list_comb (get_relmap ctxt s, args) | |
| 310 | end | |
| 311 | else | |
| 312 | let | |
| 47095 
b43ddeea727f
simplified code of generation of aggregate relations
 kuncar parents: 
46416diff
changeset | 313 | val (Type (_, rtys), qty_pat) = get_rty_qty ctxt s' | 
| 45796 | 314 | val qtyenv = match ctxt equiv_match_err qty_pat qty | 
| 47095 
b43ddeea727f
simplified code of generation of aggregate relations
 kuncar parents: 
46416diff
changeset | 315 | val rtys' = map (Envir.subst_type qtyenv) rtys | 
| 
b43ddeea727f
simplified code of generation of aggregate relations
 kuncar parents: 
46416diff
changeset | 316 | val args = map (equiv_relation ctxt) (tys ~~ rtys') | 
| 45796 | 317 | val eqv_rel = get_equiv_rel ctxt s' | 
| 318 |             val eqv_rel' = force_typ ctxt eqv_rel ([rty, rty] ---> @{typ bool})
 | |
| 319 | in | |
| 320 | if forall is_eq args | |
| 321 | then eqv_rel' | |
| 322 | else | |
| 323 | let | |
| 47095 
b43ddeea727f
simplified code of generation of aggregate relations
 kuncar parents: 
46416diff
changeset | 324 | val result = list_comb (get_relmap ctxt s, args) | 
| 45796 | 325 | in | 
| 326 | mk_rel_compose (result, eqv_rel') | |
| 327 | end | |
| 328 | end | |
| 329 | | _ => HOLogic.eq_const rty | |
| 330 | ||
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 331 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 332 | fun equiv_relation_chk ctxt (rty, qty) = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 333 | equiv_relation ctxt (rty, qty) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 334 | |> Syntax.check_term ctxt | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 335 | |
| 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 | 336 | (* 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 | 337 | |
| 47106 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 338 | exception CODE_GEN of string | 
| 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 339 | |
| 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 | 340 | 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 | 341 | 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 | 342 | 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 | 343 | in | 
| 47106 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 344 | (case Quotient_Info.lookup_quotients ctxt s of | 
| 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 345 | SOME qdata => Thm.transfer thy (#quot_thm qdata) | 
| 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 346 |     | 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 | 347 | 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 | 348 | |
| 47106 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 349 | fun get_rel_quot_thm ctxt s = | 
| 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 350 | let | 
| 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 351 | val thy = Proof_Context.theory_of ctxt | 
| 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 352 | in | 
| 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 353 | (case Quotient_Info.lookup_quotmaps ctxt s of | 
| 
dfa5ed8d94b4
use Thm.transfer for thms stored in generic context data storage
 kuncar parents: 
47096diff
changeset | 354 | 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 | 355 |     | 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 | 356 | 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 | 357 | |
| 59582 | 358 | fun is_id_quot thm = (Thm.prop_of thm = Thm.prop_of @{thm identity_quotient3})  (* FIXME equality *)
 | 
| 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 | 359 | |
| 47698 | 360 | 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 | 361 | |
| 47698 | 362 | 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 | 363 | |
| 
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 | 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 | 365 | |
| 
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 | 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 | 367 | let | 
| 59582 | 368 | val (_ $ rel $ _ $ _) = (HOLogic.dest_Trueprop o Thm.prop_of) quot_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 | 369 | 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 | 370 | 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 | 371 | 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 | 372 | |
| 
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 | 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 | 374 | 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 | 375 | 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 | 376 | in | 
| 47308 | 377 |     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 | 378 | 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 | 379 | 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 | 380 | |
| 47504 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47308diff
changeset | 381 | 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 | 382 | if rty = qty | 
| 47308 | 383 |   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 | 384 | 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 | 385 | 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 | 386 | (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 | 387 | 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 | 388 | 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 | 389 | let | 
| 47504 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47308diff
changeset | 390 | 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 | 391 | 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 | 392 | 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 | 393 | 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 | 394 | 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 | 395 | 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 | 396 | 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 | 397 | 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 | 398 | val rtys' = map (Envir.subst_type qtyenv) rtys | 
| 47504 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47308diff
changeset | 399 | 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 | 400 | 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 | 401 | 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 | 402 | 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 | 403 | 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 | 404 | 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 | 405 | 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 | 406 | 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 | 407 | 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 | 408 | 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 | 409 | 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 | 410 | 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 | 411 | end | 
| 47308 | 412 |     | _ => @{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 | 413 | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 414 | |
| 
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 | (*** Regularization ***) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 417 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 418 | (* Regularizing an rtrm means: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 419 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 420 | - Quantifiers over types that need lifting are replaced | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 421 | by bounded quantifiers, for example: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 422 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 423 | All P ----> All (Respects R) P | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 424 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 425 | 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 | 426 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 427 | - Abstractions over types that need lifting are replaced | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 428 | by bounded abstractions, for example: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 429 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 430 | %x. P ----> Ball (Respects R) %x. P | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 431 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 432 | - 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 | 433 | corresponding equivalence relations, for example: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 434 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 435 | A = B ----> R A B | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 436 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 437 | or | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 438 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 439 | A = B ----> (R ===> R) A B | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 440 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 441 | for more complicated types of A and B | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 442 | |
| 
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 | 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 | 445 | 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 | 446 | but also accepts partially regularized terms. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 447 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 448 | This means that the raw theorems can have: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 449 | 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 | 450 | in the places where: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 451 | All, Ex, Ex1, %, (op =) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 452 | is required the lifted theorem. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 453 | |
| 
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 | val mk_babs = Const (@{const_name Babs}, dummyT)
 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 457 | val mk_ball = Const (@{const_name Ball}, dummyT)
 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 458 | val mk_bex  = Const (@{const_name Bex}, dummyT)
 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 459 | 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 | 460 | val mk_resp = Const (@{const_name Respects}, dummyT)
 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 461 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 462 | (* - applies f to the subterm of an abstraction, | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 463 | otherwise to the given term, | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 464 | - used by regularize, therefore abstracted | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 465 | variables do not have to be treated specially | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 466 | *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 467 | fun apply_subt f (trm1, trm2) = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 468 | case (trm1, trm2) of | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 469 | (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 | 470 | | _ => f (trm1, trm2) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 471 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 472 | fun term_mismatch str ctxt t1 t2 = | 
| 41444 | 473 | let | 
| 474 | val t1_str = Syntax.string_of_term ctxt t1 | |
| 475 | val t2_str = Syntax.string_of_term ctxt t2 | |
| 476 | val t1_ty_str = Syntax.string_of_typ ctxt (fastype_of t1) | |
| 477 | val t2_ty_str = Syntax.string_of_typ ctxt (fastype_of t2) | |
| 478 | in | |
| 479 | raise LIFT_MATCH (cat_lines [str, t1_str ^ "::" ^ t1_ty_str, t2_str ^ "::" ^ t2_ty_str]) | |
| 480 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 481 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 482 | (* the major type of All and Ex quantifiers *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 483 | 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 | 484 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 485 | (* Checks that two types match, for example: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 486 | rty -> rty matches qty -> qty *) | 
| 45280 | 487 | 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 | 488 | let | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 489 | 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 | 490 | in | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 491 | 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 | 492 | else | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 493 | (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 | 494 | (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 | 495 | 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 | 496 | 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 | 497 | 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 | 498 | else | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 499 | (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 | 500 | 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 | 501 | | NONE => false) | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 502 | | _ => false) | 
| 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45280diff
changeset | 503 | end | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 504 | |
| 
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 | (* produces a regularized version of rtrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 507 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 508 | - the result might contain dummyTs | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 509 | |
| 38718 | 510 | - for regularization we do not need any | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 511 | special treatment of bound variables | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 512 | *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 513 | fun regularize_trm ctxt (rtrm, qtrm) = | 
| 45280 | 514 | (case (rtrm, qtrm) of | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 515 | (Abs (x, ty, t), Abs (_, ty', t')) => | 
| 41444 | 516 | let | 
| 517 | val subtrm = Abs(x, ty, regularize_trm ctxt (t, t')) | |
| 518 | in | |
| 519 | if ty = ty' then subtrm | |
| 520 | else mk_babs $ (mk_resp $ equiv_relation ctxt (ty, ty')) $ subtrm | |
| 521 | end | |
| 45280 | 522 | |
| 37677 | 523 |   | (Const (@{const_name Babs}, T) $ resrel $ (t as (Abs (_, ty, _))), t' as (Abs (_, ty', _))) =>
 | 
| 41444 | 524 | let | 
| 525 | val subtrm = regularize_trm ctxt (t, t') | |
| 526 | val needres = mk_resp $ equiv_relation_chk ctxt (ty, ty') | |
| 527 | in | |
| 528 | if resrel <> needres | |
| 529 | then term_mismatch "regularize (Babs)" ctxt resrel needres | |
| 530 | else mk_babs $ resrel $ subtrm | |
| 531 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 532 | |
| 37677 | 533 |   | (Const (@{const_name All}, ty) $ t, Const (@{const_name All}, ty') $ t') =>
 | 
| 41444 | 534 | let | 
| 535 | val subtrm = apply_subt (regularize_trm ctxt) (t, t') | |
| 536 | in | |
| 537 |         if ty = ty' then Const (@{const_name All}, ty) $ subtrm
 | |
| 538 | else mk_ball $ (mk_resp $ equiv_relation ctxt (qnt_typ ty, qnt_typ ty')) $ subtrm | |
| 539 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 540 | |
| 37677 | 541 |   | (Const (@{const_name Ex}, ty) $ t, Const (@{const_name Ex}, ty') $ t') =>
 | 
| 41444 | 542 | let | 
| 543 | val subtrm = apply_subt (regularize_trm ctxt) (t, t') | |
| 544 | in | |
| 545 |         if ty = ty' then Const (@{const_name Ex}, ty) $ subtrm
 | |
| 546 | else mk_bex $ (mk_resp $ equiv_relation ctxt (qnt_typ ty, qnt_typ ty')) $ subtrm | |
| 547 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 548 | |
| 37677 | 549 |   | (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 | 550 |       (Const (@{const_name HOL.conj}, _) $ (Const (@{const_name Set.member}, _) $ _ $
 | 
| 37677 | 551 |         (Const (@{const_name Respects}, _) $ resrel)) $ (t $ _)))),
 | 
| 552 |      Const (@{const_name Ex1}, ty') $ t') =>
 | |
| 41444 | 553 | let | 
| 554 | val t_ = incr_boundvars (~1) t | |
| 555 | val subtrm = apply_subt (regularize_trm ctxt) (t_, t') | |
| 556 | val needrel = equiv_relation_chk ctxt (qnt_typ ty, qnt_typ ty') | |
| 557 | in | |
| 558 | if resrel <> needrel | |
| 559 | then term_mismatch "regularize (Bex1)" ctxt resrel needrel | |
| 560 | else mk_bex1_rel $ resrel $ subtrm | |
| 561 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 562 | |
| 38558 | 563 |   | (Const (@{const_name Ex1}, ty) $ t, Const (@{const_name Ex1}, ty') $ t') =>
 | 
| 41444 | 564 | let | 
| 565 | val subtrm = apply_subt (regularize_trm ctxt) (t, t') | |
| 566 | in | |
| 567 |         if ty = ty' then Const (@{const_name Ex1}, ty) $ subtrm
 | |
| 568 | else mk_bex1_rel $ (equiv_relation ctxt (qnt_typ ty, qnt_typ ty')) $ subtrm | |
| 569 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 570 | |
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 571 |   | (Const (@{const_name Ball}, ty) $ (Const (@{const_name Respects}, _) $ resrel) $ t,
 | 
| 38558 | 572 |      Const (@{const_name All}, ty') $ t') =>
 | 
| 41444 | 573 | let | 
| 574 | val subtrm = apply_subt (regularize_trm ctxt) (t, t') | |
| 575 | val needrel = equiv_relation_chk ctxt (qnt_typ ty, qnt_typ ty') | |
| 576 | in | |
| 577 | if resrel <> needrel | |
| 578 | then term_mismatch "regularize (Ball)" ctxt resrel needrel | |
| 579 | else mk_ball $ (mk_resp $ resrel) $ subtrm | |
| 580 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 581 | |
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 582 |   | (Const (@{const_name Bex}, ty) $ (Const (@{const_name Respects}, _) $ resrel) $ t,
 | 
| 38558 | 583 |      Const (@{const_name Ex}, ty') $ t') =>
 | 
| 41444 | 584 | let | 
| 585 | val subtrm = apply_subt (regularize_trm ctxt) (t, t') | |
| 586 | val needrel = equiv_relation_chk ctxt (qnt_typ ty, qnt_typ ty') | |
| 587 | in | |
| 588 | if resrel <> needrel | |
| 589 | then term_mismatch "regularize (Bex)" ctxt resrel needrel | |
| 590 | else mk_bex $ (mk_resp $ resrel) $ subtrm | |
| 591 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 592 | |
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 593 |   | (Const (@{const_name Bex1_rel}, ty) $ resrel $ t, Const (@{const_name Ex1}, ty') $ t') =>
 | 
| 41444 | 594 | let | 
| 595 | val subtrm = apply_subt (regularize_trm ctxt) (t, t') | |
| 596 | val needrel = equiv_relation_chk ctxt (qnt_typ ty, qnt_typ ty') | |
| 597 | in | |
| 598 | if resrel <> needrel | |
| 599 | then term_mismatch "regularize (Bex1_res)" ctxt resrel needrel | |
| 600 | else mk_bex1_rel $ resrel $ subtrm | |
| 601 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 602 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 603 | | (* equalities need to be replaced by appropriate equivalence relations *) | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 604 |     (Const (@{const_name HOL.eq}, ty), Const (@{const_name HOL.eq}, ty')) =>
 | 
| 41444 | 605 | if ty = ty' then rtrm | 
| 606 | 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 | 607 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 608 | | (* 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 | 609 |     (rel, Const (@{const_name HOL.eq}, ty')) =>
 | 
| 41444 | 610 | let | 
| 611 | val rel_ty = fastype_of rel | |
| 612 | val rel' = equiv_relation_chk ctxt (domain_type rel_ty, domain_type ty') | |
| 613 | in | |
| 614 | if rel' aconv rel then rtrm | |
| 615 | else term_mismatch "regularize (relation mismatch)" ctxt rel rel' | |
| 616 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 617 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 618 | | (_, Const _) => | 
| 41444 | 619 | let | 
| 42361 | 620 | val thy = Proof_Context.theory_of ctxt | 
| 45280 | 621 | fun same_const (Const (s, T)) (Const (s', T')) = s = s' andalso matches_typ ctxt T T' | 
| 41444 | 622 | | same_const _ _ = false | 
| 623 | in | |
| 624 | if same_const rtrm qtrm then rtrm | |
| 625 | else | |
| 626 | let | |
| 45279 | 627 | 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 | 628 | (case Quotient_Info.lookup_quotconsts_global thy qtrm of | 
| 45279 | 629 | SOME qconst_info => #rconst qconst_info | 
| 630 | | NONE => term_mismatch "regularize (constant not found)" ctxt rtrm qtrm) | |
| 41444 | 631 | in | 
| 632 | if Pattern.matches thy (rtrm', rtrm) | |
| 633 | then rtrm else term_mismatch "regularize (constant mismatch)" ctxt rtrm qtrm | |
| 634 | end | |
| 635 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 636 | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
47698diff
changeset | 637 |   | (((t1 as Const (@{const_name case_prod}, _)) $ Abs (v1, ty, Abs(v1', ty', s1))),
 | 
| 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
47698diff
changeset | 638 |      ((t2 as Const (@{const_name case_prod}, _)) $ Abs (v2, _ , Abs(v2', _  , s2)))) =>
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 639 | 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 | 640 | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
47698diff
changeset | 641 |   | (((t1 as Const (@{const_name case_prod}, _)) $ Abs (v1, ty, s1)),
 | 
| 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
47698diff
changeset | 642 |      ((t2 as Const (@{const_name case_prod}, _)) $ Abs (v2, _ , s2))) =>
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 643 | 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 | 644 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 645 | | (t1 $ t2, t1' $ t2') => | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 646 | 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 | 647 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 648 | | (Bound i, Bound i') => | 
| 41444 | 649 | if i = i' then rtrm | 
| 650 | else raise (LIFT_MATCH "regularize (bounds mismatch)") | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 651 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 652 | | _ => | 
| 41444 | 653 | let | 
| 654 | val rtrm_str = Syntax.string_of_term ctxt rtrm | |
| 655 | val qtrm_str = Syntax.string_of_term ctxt qtrm | |
| 656 | in | |
| 657 |         raise (LIFT_MATCH ("regularize failed (default: " ^ rtrm_str ^ "," ^ qtrm_str ^ ")"))
 | |
| 45280 | 658 | end) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 659 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 660 | fun regularize_trm_chk ctxt (rtrm, qtrm) = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 661 | regularize_trm ctxt (rtrm, qtrm) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 662 | |> Syntax.check_term ctxt | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 663 | |
| 
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 | (*** Rep/Abs Injection ***) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 667 | |
| 
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 | Injection of Rep/Abs means: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 670 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 671 | For abstractions: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 672 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 673 | * 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 | 674 | around the abstraction; otherwise we leave it unchanged. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 675 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 676 | For applications: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 677 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 678 | * 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 | 679 | 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 | 680 | 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 | 681 | 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 | 682 | arguments. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 683 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 684 | For constants: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 685 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 686 | * 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 | 687 | 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 | 688 | and Rep/Abs around it. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 689 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 690 | For free variables: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 691 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 692 | * 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 | 693 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 694 | Vars case cannot occur. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 695 | *) | 
| 
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 | fun mk_repabs ctxt (T, T') trm = | 
| 45797 | 698 | 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 | 699 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 700 | fun inj_repabs_err ctxt msg rtrm qtrm = | 
| 41444 | 701 | let | 
| 702 | val rtrm_str = Syntax.string_of_term ctxt rtrm | |
| 703 | val qtrm_str = Syntax.string_of_term ctxt qtrm | |
| 704 | in | |
| 705 | raise LIFT_MATCH (space_implode " " [msg, quote rtrm_str, "and", quote qtrm_str]) | |
| 706 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 707 | |
| 
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 | (* bound variables need to be treated properly, | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 710 | 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 | 711 | fun inj_repabs_trm ctxt (rtrm, qtrm) = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 712 | case (rtrm, qtrm) of | 
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 713 |     (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 | 714 |        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 | 715 | |
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 716 |   | (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 | 717 |        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 | 718 | |
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 719 |   | (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 | 720 | let | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 721 | val rty = fastype_of rtrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 722 | val qty = fastype_of qtrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 723 | in | 
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 724 |         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 | 725 | end | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 726 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 727 | | (Abs (x, T, t), Abs (x', T', t')) => | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 728 | let | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 729 | val rty = fastype_of rtrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 730 | val qty = fastype_of qtrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 731 | 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 | 732 | val (_, 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 yvar = Free (y, T) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 734 | 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 | 735 | in | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 736 | if rty = qty then result | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 737 | else mk_repabs ctxt (rty, qty) result | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 738 | end | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 739 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 740 | | (t $ s, t' $ s') => | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 741 | (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 | 742 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 743 | | (Free (_, T), Free (_, T')) => | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 744 | if T = T' then rtrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 745 | else mk_repabs ctxt (T, T') rtrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 746 | |
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 747 |   | (_, Const (@{const_name HOL.eq}, _)) => rtrm
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 748 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 749 | | (_, Const (_, T')) => | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 750 | let | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 751 | val rty = fastype_of rtrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 752 | in | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 753 | if rty = T' then rtrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 754 | else mk_repabs ctxt (rty, T') rtrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 755 | end | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 756 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 757 | | _ => inj_repabs_err ctxt "injection (default):" rtrm qtrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 758 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 759 | fun inj_repabs_trm_chk ctxt (rtrm, qtrm) = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 760 | inj_repabs_trm ctxt (rtrm, qtrm) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 761 | |> Syntax.check_term ctxt | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 762 | |
| 
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 | (*** 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 | 766 | |
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 767 | (* 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 | 768 | 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 | 769 | *) | 
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 770 | 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 | 771 | case rty of | 
| 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 772 | Type (s, rtys) => | 
| 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 773 | let | 
| 42361 | 774 | 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 | 775 | 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 | 776 | |
| 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 777 | fun matches [] = rty' | 
| 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 778 | | matches ((rty, qty)::tail) = | 
| 45280 | 779 | (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 | 780 | 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 | 781 | | 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 | 782 | in | 
| 41444 | 783 | matches ty_subst | 
| 784 | end | |
| 37560 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 785 | | _ => rty | 
| 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 786 | |
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 787 | 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 | 788 | 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 | 789 | 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 | 790 | | 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 | 791 | | 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 | 792 | | 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 | 793 | | Bound i => Bound i | 
| 41444 | 794 | | Const (a, ty) => | 
| 37560 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 795 | let | 
| 42361 | 796 | val thy = Proof_Context.theory_of ctxt | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 797 | |
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 798 | 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 | 799 | | matches ((rconst, qconst)::tail) = | 
| 45280 | 800 | (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 | 801 | 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 | 802 | | 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 | 803 | in | 
| 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 804 | matches trm_subst | 
| 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 805 | end | 
| 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 806 | |
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 807 | (* generate type and term substitutions out of the | 
| 41444 | 808 | qtypes involved in a quotient; the direction flag | 
| 809 | indicates in which direction the substitutions work: | |
| 810 | ||
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 811 | true: quotient -> raw | 
| 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 812 | false: raw -> quotient | 
| 37560 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 813 | *) | 
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 814 | fun mk_ty_subst qtys direction ctxt = | 
| 41444 | 815 | let | 
| 42361 | 816 | val thy = Proof_Context.theory_of ctxt | 
| 41444 | 817 | in | 
| 45279 | 818 | Quotient_Info.dest_quotients ctxt | 
| 41444 | 819 | |> map (fn x => (#rtyp x, #qtyp x)) | 
| 820 | |> filter (fn (_, qty) => member (Sign.typ_instance thy o swap) qtys qty) | |
| 821 | |> map (if direction then swap else I) | |
| 822 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 823 | |
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 824 | fun mk_trm_subst qtys direction ctxt = | 
| 41444 | 825 | let | 
| 826 | val subst_typ' = subst_typ ctxt (mk_ty_subst qtys direction ctxt) | |
| 827 | 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 | 828 | |
| 41444 | 829 | val const_substs = | 
| 45279 | 830 | Quotient_Info.dest_quotconsts ctxt | 
| 41444 | 831 | |> map (fn x => (#rconst x, #qconst x)) | 
| 832 | |> 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 | 833 | |
| 41444 | 834 | val rel_substs = | 
| 45279 | 835 | Quotient_Info.dest_quotients ctxt | 
| 41444 | 836 | |> map (fn x => (#equiv_rel x, HOLogic.eq_const (#qtyp x))) | 
| 837 | |> map (if direction then swap else I) | |
| 838 | in | |
| 839 | filter proper (const_substs @ rel_substs) | |
| 840 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 841 | |
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 842 | |
| 37560 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 843 | (* derives a qtyp and qtrm out of a rtyp and rtrm, | 
| 41444 | 844 | respectively | 
| 37560 
1b5bbc4a14bc
streamlined the generation of quotient theorems out of raw theorems
 Christian Urban <urbanc@in.tum.de> parents: 
37532diff
changeset | 845 | *) | 
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 846 | 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 | 847 | 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 | 848 | |
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 849 | 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 | 850 | 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 | 851 | |
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 852 | (* derives a rtyp and rtrm out of a qtyp and qtrm, | 
| 41444 | 853 | respectively | 
| 37592 
e16495cfdde0
separation of translations in derive_qtrm / derive_rtrm (similarly for types)
 Christian Urban <urbanc@in.tum.de> parents: 
37591diff
changeset | 854 | *) | 
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 855 | 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 | 856 | 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 | 857 | |
| 38624 
9bb0016f7e60
changed to a more convenient argument order
 Christian Urban <urbanc@in.tum.de> parents: 
38558diff
changeset | 858 | 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 | 859 | 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 | 860 | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 861 | |
| 45279 | 862 | end; |