| author | wenzelm | 
| Thu, 26 Sep 2024 21:55:46 +0200 | |
| changeset 80966 | cf96b584724d | 
| parent 80636 | 4041e7c8059d | 
| child 81807 | cd2521e39783 | 
| permissions | -rw-r--r-- | 
| 60918 | 1 | (* Title: HOL/Tools/BNF/bnf_lift.ML | 
| 2 | Author: Julian Biendarra, TU Muenchen | |
| 71262 | 3 | Author: Basil Fürer, ETH Zurich | 
| 4 | Author: Joshua Schneider, ETH Zurich | |
| 60918 | 5 | Author: Dmitriy Traytel, ETH Zurich | 
| 71262 | 6 | Copyright 2015, 2019 | 
| 60918 | 7 | |
| 8 | Lifting of BNFs through typedefs. | |
| 9 | *) | |
| 10 | ||
| 61067 | 11 | signature BNF_LIFT = | 
| 12 | sig | |
| 13 | datatype lift_bnf_option = | |
| 14 | Plugins_Option of Proof.context -> Plugin_Name.filter | |
| 15 | | No_Warn_Wits | |
| 71262 | 16 | | No_Warn_Transfer | 
| 60918 | 17 | val copy_bnf: | 
| 18 | (((lift_bnf_option list * (binding option * (string * sort option)) list) * | |
| 62324 | 19 | string) * thm option) * (binding * binding * binding) -> | 
| 60918 | 20 | local_theory -> local_theory | 
| 21 | val copy_bnf_cmd: | |
| 22 | (((lift_bnf_option list * (binding option * (string * string option)) list) * | |
| 62324 | 23 | string) * (Facts.ref * Token.src list) option) * (binding * binding * binding) -> | 
| 60918 | 24 | local_theory -> local_theory | 
| 25 | val lift_bnf: | |
| 62777 | 26 | ((((lift_bnf_option list * (binding option * (string * sort option)) list) * | 
| 71494 | 27 | string) * term list option) * thm list option) * (binding * binding * binding) -> | 
| 60918 | 28 |       ({context: Proof.context, prems: thm list} -> tactic) list ->
 | 
| 29 | local_theory -> local_theory | |
| 30 | val lift_bnf_cmd: | |
| 31 | ((((lift_bnf_option list * (binding option * (string * string option)) list) * | |
| 71494 | 32 | string) * string list) * (Facts.ref * Token.src list) list option) * | 
| 62324 | 33 | (binding * binding * binding) -> local_theory -> Proof.state | 
| 61067 | 34 | end | 
| 60918 | 35 | |
| 61067 | 36 | structure BNF_Lift : BNF_LIFT = | 
| 37 | struct | |
| 60918 | 38 | |
| 39 | open Ctr_Sugar_Tactics | |
| 40 | open BNF_Util | |
| 41 | open BNF_Comp | |
| 42 | open BNF_Def | |
| 43 | ||
| 61067 | 44 | |
| 45 | datatype lift_bnf_option = | |
| 46 | Plugins_Option of Proof.context -> Plugin_Name.filter | |
| 71262 | 47 | | No_Warn_Wits | 
| 48 | | No_Warn_Transfer; | |
| 49 | ||
| 71494 | 50 | datatype equiv_thm = Typedef | Quotient of thm | 
| 71262 | 51 | |
| 52 | (** Util **) | |
| 53 | ||
| 54 | fun last2 [x, y] = ([], (x, y)) | |
| 55 | | last2 (x :: xs) = last2 xs |>> (fn y => x :: y) | |
| 56 | | last2 [] = raise Match; | |
| 57 | ||
| 58 | fun strip3 thm = (case Term.strip_comb (HOLogic.dest_Trueprop (Thm.prop_of thm)) of | |
| 59 | (_, [x1, x2, x3]) => (x1, x2, x3) | |
| 60 | | _ => error "strip3: wrong number of arguments"); | |
| 61 | ||
| 62 | val mk_Free = yield_singleton o mk_Frees; | |
| 63 | ||
| 64 | fun TWICE t = t THEN' t; | |
| 65 | ||
| 66 | fun prove lthy fvars tm tac = | |
| 67 |   Goal.prove_sorry lthy (map (fst o dest_Free) fvars) [] tm (fn {context, ...} => tac context);
 | |
| 68 | ||
| 69 | (** Term construction **) | |
| 70 | ||
| 71 | fun mk_relT aT bT = aT --> bT --> HOLogic.boolT; | |
| 72 | fun mk_relcompp r s = let | |
| 73 | val (rT, sT) = apply2 fastype_of (r, s); | |
| 74 | val ((xT, _), (_, zTs)) = apply2 dest_funT (rT, sT); | |
| 75 | val T = rT --> sT --> mk_relT xT (fst (dest_funT zTs)); | |
| 76 |   in Const (@{const_name relcompp}, T) $ r $ s end;
 | |
| 77 | val mk_OO = fold_rev mk_relcompp; | |
| 78 | ||
| 79 | (* option from sum *) | |
| 80 | fun mk_MaybeT T = mk_sumT (HOLogic.unitT, T); | |
| 81 | fun mk_Nothing T = BNF_FP_Util.mk_Inl T HOLogic.unit; | |
| 82 | val Just_const = BNF_FP_Util.Inr_const HOLogic.unitT; | |
| 83 | fun mk_Just tm = Just_const (fastype_of tm) $ tm; | |
| 84 | fun Maybe_map_const T = | |
| 85 | let val (xT, yT) = dest_funT T in | |
| 86 |     Const (@{const_name map_sum}, (HOLogic.unitT --> HOLogic.unitT) --> T --> mk_MaybeT xT --> mk_MaybeT yT) $
 | |
| 87 | HOLogic.id_const HOLogic.unitT | |
| 88 | end; | |
| 89 | fun mk_Maybe_map tm = Maybe_map_const (fastype_of tm) $ tm; | |
| 90 | fun fromJust_const T = Const (@{const_name sum.projr}, mk_MaybeT T --> T)
 | |
| 91 | ||
| 92 | fun rel_Maybe_const T U = | |
| 93 |   Const (@{const_name rel_sum}, (HOLogic.unitT --> HOLogic.unitT --> HOLogic.boolT) -->
 | |
| 94 | (T --> U --> HOLogic.boolT) --> mk_MaybeT T --> mk_MaybeT U --> HOLogic.boolT) $ | |
| 95 | HOLogic.eq_const HOLogic.unitT | |
| 96 | fun set_Maybe_const T = Const (@{const_name Basic_BNFs.setr}, mk_MaybeT T --> HOLogic.mk_setT T)
 | |
| 97 | ||
| 98 | fun Inf_const T = Const (@{const_name Inf}, HOLogic.mk_setT T --> T);
 | |
| 99 | ||
| 100 | fun Image_const T = | |
| 101 | let | |
| 102 | val relT = HOLogic.mk_setT (HOLogic.mk_prodT (T, T)); | |
| 103 | val setT = HOLogic.mk_setT T; | |
| 104 |   in Const (@{const_name Image}, relT --> setT --> setT) end;
 | |
| 105 | ||
| 106 | fun bot_const T = Const (@{const_name bot}, T);
 | |
| 107 | ||
| 108 | fun mk_insert x S = | |
| 109 |   Const (@{const_name Set.insert}, fastype_of x --> fastype_of S --> fastype_of S) $ x $ S;
 | |
| 110 | ||
| 111 | fun mk_vimage f s = | |
| 112 | let val (xT, yT) = dest_funT (fastype_of f) in | |
| 113 |     Const (@{const_name vimage}, (xT --> yT) --> HOLogic.mk_setT yT --> HOLogic.mk_setT xT) $ f $ s
 | |
| 114 | end; | |
| 115 | ||
| 116 | fun mk_case_prod (x, y) tm = let | |
| 117 | val ((x, xT), (y, yT)) = apply2 dest_Free (x, y); | |
| 118 | val prodT = HOLogic.mk_prodT (xT, yT); | |
| 119 |    in HOLogic.Collect_const prodT $ (Const (@{const_name case_prod},
 | |
| 120 | (xT --> yT --> HOLogic.boolT) --> prodT --> HOLogic.boolT) $ absfree (x, xT) | |
| 121 | (absfree (y, yT) tm)) end; | |
| 122 | ||
| 123 | fun mk_Trueprop_implies (ps, c) = | |
| 124 | Logic.list_implies (map HOLogic.mk_Trueprop ps, HOLogic.mk_Trueprop c); | |
| 125 | ||
| 126 | fun mk_Collect (v, tm) = let val (n, T) = dest_Free v in | |
| 127 | HOLogic.mk_Collect (n, T, tm) end; | |
| 128 | ||
| 129 | ||
| 130 | (** witnesses **) | |
| 131 | fun prepare_wits is_quotient RepT wits opts alphas wits_F var_as var_as' sets lthy = | |
| 132 | let | |
| 133 | fun binder_types_until_eq V T = | |
| 134 | let | |
| 135 |         fun strip (TU as Type ("fun", [T, U])) = if V = TU then [] else T :: strip U
 | |
| 136 | | strip T = if V = T then [] else | |
| 137 |               error ("Bad type for witness: " ^ quote (Syntax.string_of_typ lthy T));
 | |
| 138 | in strip T end; | |
| 139 | ||
| 140 | val Iwits = the_default wits_F (Option.map (map (`(map (fn T => | |
| 141 | find_index (fn U => T = U) alphas) o binder_types_until_eq RepT o fastype_of))) wits); | |
| 142 | ||
| 71469 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 143 | val Iwits = if is_quotient then | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 144 | let | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 145 | val subsumed_Iwits = | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 146 | filter (fn (J, _) => exists (fn (I, _) => subset (op =) (I, J)) wits_F) Iwits; | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 147 | val _ = if null subsumed_Iwits orelse exists (fn No_Warn_Wits => true | _ => false) opts | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 148 | then () | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 149 | else | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 150 | let | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 151 | val (suffix1, suffix2, be) = | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 152 |                   (if length subsumed_Iwits = 1 then ("", "", "is") else ("s", "es", "are"))
 | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 153 | in | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 154 | subsumed_Iwits | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 155 | |> map (Syntax.pretty_typ lthy o fastype_of o snd) | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 156 | |> Pretty.big_list | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 157 |                   ("The following type" ^ suffix1 ^ " of nonemptiness witness" ^ suffix2 ^
 | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 158 | " of the raw type's BNF " ^ be ^ " subsumed by the existing raw witnesses:") | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 159 | |> (fn pt => Pretty.chunks [pt, | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 160 |                   Pretty.para ("You do not need to lift these subsumed witnesses.")])
 | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 161 | |> Pretty.string_of | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 162 | |> warning | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 163 | end; | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 164 | in | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 165 | filter_out (fn (J, _) => exists (fn (I, _) => subset (op =) (I, J)) wits_F) Iwits | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 166 | end | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 167 | else Iwits; | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 168 | |
| 71262 | 169 | val wit_goals = maps (BNF_Def.mk_wit_goals var_as var_as' sets) Iwits; | 
| 170 | ||
| 71469 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 171 | val lost_wits = if is_quotient then [] else | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 172 | filter_out (fn (J, _) => exists (fn (I, _) => subset (op =) (I, J)) Iwits) wits_F; | 
| 71262 | 173 | |
| 174 | val _ = | |
| 175 | if null lost_wits orelse exists (fn No_Warn_Wits => true | _ => false) opts then () | |
| 176 | else | |
| 177 | let | |
| 178 | val what = (if is_quotient then "quotient type" else "typedef"); | |
| 179 | val (suffix1, suffix2, be) = | |
| 180 |             (if length lost_wits = 1 then ("", "", "was") else ("s", "es", "were"))
 | |
| 181 | in | |
| 182 | lost_wits | |
| 183 | |> map (Syntax.pretty_typ lthy o fastype_of o snd) | |
| 184 | |> Pretty.big_list | |
| 185 |             ("The following type" ^ suffix1 ^ " of nonemptiness witness" ^ suffix2 ^
 | |
| 186 | " of the raw type's BNF " ^ be ^ " lost:") | |
| 187 | |> (fn pt => Pretty.chunks [pt, | |
| 188 |             Pretty.para ("You can specify a liftable witness (e.g., a term of one of the above types\
 | |
| 189 | \ that satisfies the " ^ what ^ "'s invariant)\ | |
| 190 | \ using the annotation [wits: <term>].")]) | |
| 191 | |> Pretty.string_of | |
| 192 | |> warning | |
| 193 | end; | |
| 194 | in (Iwits, wit_goals) end; | |
| 195 | ||
| 196 | ||
| 197 | (** transfer theorems **) | |
| 198 | ||
| 199 | fun mk_transfer_thms' bnf_F bnf_G name consts Tss crel_def pcrel_def defs lthy = let | |
| 200 | ||
| 201 | val live = length (#alphas Tss); | |
| 202 | ||
| 203 | val (pcrel_tm, crel_tm) = apply2 (Thm.prop_of #> Logic.dest_equals #> fst #> head_of) | |
| 204 | (pcrel_def, crel_def); | |
| 205 | ||
| 79564 
33b10cd883ae
made lift_bnf more robust for abstract types with 'phantom' type variables
 traytel parents: 
75624diff
changeset | 206 | val (var_Qs, var_Rs) = lthy | 
| 
33b10cd883ae
made lift_bnf more robust for abstract types with 'phantom' type variables
 traytel parents: 
75624diff
changeset | 207 | |> fold Variable.declare_typ (#alphas Tss @ Library.union (op =) (#deads Tss) (#Ds0 Tss)) | 
| 71262 | 208 | |> mk_Frees "Q" (map2 mk_relT (#alphas Tss) (#betas Tss)) | 
| 209 | ||>> mk_Frees "R" (map2 mk_relT (#gammas Tss) (#deltas Tss)) | |
| 210 | |> fst; | |
| 211 | ||
| 212 | (* get the "pcrel :: args_raw => rep => abs \<Rightarrow> bool" term and instantiate types *) | |
| 213 | val (args_raw, (rep, abs)) = pcrel_tm | |
| 214 | |> fastype_of | |
| 215 | |> binder_types | |
| 216 | |> last2; | |
| 217 | val thy = Proof_Context.theory_of lthy; | |
| 218 | val tyenv_match = Vartab.empty | |
| 219 | |> Sign.typ_match thy ((rep, #rep Tss)) | |
| 220 | |> Sign.typ_match thy ((abs, #abs Tss)); | |
| 221 | val args = map (Envir.subst_type tyenv_match) args_raw; | |
| 222 | val (pcrel_a, pcrel_b) = Envir.subst_term (tyenv_match, Vartab.empty) pcrel_tm | |
| 223 | |> `(subst_atomic_types (#alphas Tss @ #betas Tss ~~ #gammas Tss @ #deltas Tss)); | |
| 224 | ||
| 79564 
33b10cd883ae
made lift_bnf more robust for abstract types with 'phantom' type variables
 traytel parents: 
75624diff
changeset | 225 |     (* match "crel :: ?a F \<Rightarrow> {a G} \<Rightarrow> bool" with "a G" *)
 | 
| 71262 | 226 | val tyenv_match = Vartab.empty |> Sign.typ_match thy | 
| 79564 
33b10cd883ae
made lift_bnf more robust for abstract types with 'phantom' type variables
 traytel parents: 
75624diff
changeset | 227 | (crel_tm |> fastype_of |> binder_types |> tl |> hd, #abs Tss); | 
| 71262 | 228 | val crel_b = Envir.subst_term (tyenv_match, Vartab.empty) crel_tm | 
| 79564 
33b10cd883ae
made lift_bnf more robust for abstract types with 'phantom' type variables
 traytel parents: 
75624diff
changeset | 229 | |> subst_atomic_types (#alphas Tss ~~ #betas Tss); | 
| 71262 | 230 | val crel_d = subst_atomic_types (#betas Tss ~~ #deltas Tss) crel_b; | 
| 231 | ||
| 232 | (* instantiate pcrel with Qs and Rs*) | |
| 233 | val dead_args = map binder_types args | |
| 234 | |> map (fn [T,U] => if T = U then SOME T else NONE | _ => NONE); | |
| 235 | val parametrize = let | |
| 236 | fun go (SOME T :: dTs) tms = HOLogic.eq_const T :: go dTs tms | |
| 237 | | go (_ :: dTs) (tm :: tms) = tm :: go dTs tms | |
| 238 | | go (_ :: dTs) tms = go dTs tms | |
| 239 | | go _ _ = []; | |
| 240 | in go dead_args end; | |
| 241 | val pcrel_Qs = list_comb (pcrel_b, parametrize var_Qs); | |
| 242 | val pcrel_Rs = list_comb (pcrel_a, parametrize var_Rs); | |
| 243 | ||
| 244 | (* get the order of the dead variables right *) | |
| 79564 
33b10cd883ae
made lift_bnf more robust for abstract types with 'phantom' type variables
 traytel parents: 
75624diff
changeset | 245 | val Ds0 = Library.union (op =) (maps the_list dead_args) (#Ds0 Tss); | 
| 71262 | 246 | val permute_Ds = (mk_T_of_bnf Ds0 (#betas Tss) bnf_G, nth (binder_types (type_of pcrel_Qs)) 1) | 
| 247 | |> apply2 (fn Type (_, Ts) => Ts | _ => []) |> op~~ |> typ_subst_atomic; | |
| 248 | val Ds0 = map permute_Ds Ds0; | |
| 249 | ||
| 250 | (* terms for sets of the set_transfer thms *) | |
| 251 |     val rel_sets_Q = @{map 3} (fn aT => fn bT => fn Q =>
 | |
| 252 |       mk_rel 1 [aT] [bT] @{term rel_set} $ Q) (#alphas Tss) (#betas Tss) var_Qs;
 | |
| 253 | ||
| 254 | (* rewrite rules for pcrel and BNF's relator: "pcrel Q = rel_F OO crel" *) | |
| 255 | fun mk_pcr_rel_F_eq Ts Us pcrel vs crel = | |
| 256 | let | |
| 257 | val thm = HOLogic.mk_Trueprop (HOLogic.mk_eq (pcrel, mk_relcompp (list_comb | |
| 258 | (mk_rel_of_bnf (#deads Tss) (Ts Tss) (Us Tss) bnf_F, vs)) crel)); | |
| 259 |         fun tac ctxt = unfold_thms_tac ctxt (pcrel_def :: defs @ @{thm id_bnf_apply} ::
 | |
| 260 | Transfer.get_relator_eq ctxt) THEN (HEADGOAL (rtac ctxt refl)) | |
| 261 | in prove lthy vs thm tac |> mk_abs_def end; | |
| 262 | ||
| 263 | val pcr_rel_F_eqs = | |
| 264 | [mk_pcr_rel_F_eq #alphas #betas pcrel_Qs var_Qs crel_b, | |
| 265 | mk_pcr_rel_F_eq #gammas #deltas pcrel_Rs var_Rs crel_d]; | |
| 266 | ||
| 267 |     (* create transfer-relations from term('s type) *)
 | |
| 268 | fun mk_transfer_rel' i tm = | |
| 269 | let | |
| 270 | fun go T' (n, q) = case T' of | |
| 271 |             Type ("fun", [T as Type ("fun", _), U]) =>
 | |
| 272 | go U (n+1, q) |>> mk_rel_fun (fst (go T (n, q))) | |
| 273 |           | Type ("fun", [T, U]) =>
 | |
| 274 | go T (n, q) |-> (fn x => fn st => go U st |>> mk_rel_fun x) | |
| 275 |           | Type (@{type_name bool}, _) => (HOLogic.eq_const HOLogic.boolT, (n, q))
 | |
| 276 |           | Type (@{type_name set}, _) => (nth rel_sets_Q n, (n, q))
 | |
| 277 | | Type _ => (if q then pcrel_Qs else pcrel_Rs, (n, not q)) | |
| 278 | | TFree _ => (nth (if q then var_Qs else var_Rs) n, (n, not q)) | |
| 279 | | _ => raise Match | |
| 280 | in go (fastype_of tm) (i, true) |> fst end; | |
| 281 | ||
| 282 | (* prove transfer rules *) | |
| 283 | fun prove_transfer_thm' i vars new_tm const = | |
| 284 | let | |
| 285 | val tm = HOLogic.mk_Trueprop (mk_transfer_rel' i new_tm $ #raw const $ new_tm); | |
| 286 | val tac = (fn ctxt => unfold_thms_tac ctxt (pcr_rel_F_eqs @ [crel_def]) THEN | |
| 287 |           #tac const {Rs=var_Rs,Qs=var_Qs} ctxt);
 | |
| 288 | in prove lthy vars tm tac end; | |
| 289 | val prove_transfer_thm = prove_transfer_thm' 0; | |
| 290 | (* map transfer: "((Q ===> R) ===> pcr_G Q ===> pcr_G R) map_F' map_G" *) | |
| 291 | val map_G = mk_map_of_bnf Ds0 (#betas Tss) (#deltas Tss) bnf_G; | |
| 292 | val map_transfer = prove_transfer_thm (var_Qs @ var_Rs) map_G (#map consts); | |
| 293 | (* pred_transfer: "((Q ===> (=)) ===> pcr_G Q ===> (=)) pred_F' pred_G" *) | |
| 294 | val pred_G = mk_pred_of_bnf Ds0 (#betas Tss) bnf_G; | |
| 295 | val pred_transfer = #pred consts |> Option.map (fn p => | |
| 296 |       ("pred", [prove_transfer_thm (var_Qs @ var_Rs) pred_G p]));
 | |
| 297 | (* rel_transfer: "((Q ===> R ===> (=)) ===> pcr_G Q ===> pcr_G R ===> (=)) rel_F' rel_G" *) | |
| 298 | val rel_G = mk_rel_of_bnf Ds0 (#betas Tss) (#deltas Tss) bnf_G; | |
| 299 | val rel_transfer = prove_transfer_thm (var_Qs @ var_Rs) rel_G (#rel consts); | |
| 300 | (* set_transfer: "(pcr_G Q ===> rel_set Q) set_F' set_G" *) | |
| 301 | val sets_G = mk_sets_of_bnf (replicate live Ds0) (replicate live (#betas Tss)) bnf_G; | |
| 302 |     fun mk_set_transfer i set_G raw tac = prove_transfer_thm' i var_Qs set_G {raw=raw, tac=tac};
 | |
| 303 |     val sets_transfer = @{map 4} mk_set_transfer
 | |
| 304 | (0 upto (live-1)) sets_G (#raws (#sets consts)) (#tacs (#sets consts)); | |
| 305 | (* export transfer theorems *) | |
| 306 | val transform = Morphism.thm (Morphism.thm_morphism "BNF" (unfold_thms lthy defs)) |> map; | |
| 307 | val b = Binding.qualified_name name; | |
| 308 | val qualify = | |
| 309 | let val qs = Binding.path_of b; | |
| 310 | in fold_rev (fn (s, mand) => Binding.qualify mand s) qs end; | |
| 311 | fun mk_binding n = Binding.name (n ^ "_transfer_raw") | |
| 312 | |> Binding.qualify true (Binding.name_of b) |> qualify; | |
| 313 |     val notes = [("map", [map_transfer]), ("rel", [rel_transfer])] @ the_list pred_transfer @
 | |
| 314 |       [("set", sets_transfer)] |> map (fn (n, thms) =>
 | |
| 315 |         ((mk_binding n, []), [(transform thms, @{attributes [transfer_rule]})]));
 | |
| 316 | ||
| 317 | in lthy |> Local_Theory.notes notes |> snd end; | |
| 318 | ||
| 319 | (* transfer theorems for map, pred (in case of a typedef), rel and sets *) | |
| 320 | fun mk_transfer_thms quiet bnf_F bnf_G name consts thm Tss defs lthy = let | |
| 321 | ||
| 322 | fun mk_crel_def quot_thm = | |
| 323 | (case thm of | |
| 324 |         Quotient equiv => @{thm Quotient_crel_quotient} OF [quot_thm, equiv]
 | |
| 71494 | 325 |       | Typedef => hd ([quot_thm] RL @{thms Quotient_crel_typedef Quotient_crel_typecopy}));
 | 
| 71262 | 326 |     fun no_quotient _ = [Pretty.para ("No quotient theorem has been registered for " ^
 | 
| 327 | Binding.name_of (name_of_bnf bnf_G) ^ "."), | |
| 328 | Pretty.para "Use setup_lifting to register a quotient or type definition theorem."]; | |
| 329 |     fun wrong_quotient T lthy = [Pretty.para ("A wrong quotient theorem has been registered for " ^
 | |
| 330 | Binding.name_of (name_of_bnf bnf_G) ^ "."), | |
| 331 |       Pretty.para ("Expected raw type " ^
 | |
| 332 | Pretty.string_of (Syntax.pretty_typ lthy (T_of_bnf bnf_F)) ^ | |
| 333 | " but the quotient theorem has raw type " ^ | |
| 334 | Pretty.string_of (Syntax.pretty_typ lthy T) ^ "."), | |
| 335 | Pretty.para "Use setup_lifting to register a different quotient or type definition theorem."]; | |
| 336 |     fun pcr_why _ = [Pretty.para ("The pcr_" ^ Binding.name_of (name_of_bnf bnf_G) ^
 | |
| 337 | " relator has not been defined.")]; | |
| 338 | fun warn_transfer why lthy = | |
| 339 | (Pretty.para "The transfer theorems can't be generated:" :: why lthy) | |
| 340 | |> Pretty.chunks |> Pretty.string_of |> warning |> K lthy; | |
| 341 | fun maybe_warn_transfer why = not quiet ? warn_transfer why; | |
| 342 | in | |
| 343 | case Lifting_Info.lookup_quotients lthy name of | |
| 344 |       SOME {pcr_info, quot_thm} =>
 | |
| 345 | (let | |
| 346 | val crel_def = mk_crel_def quot_thm; | |
| 347 | val rty = Lifting_Util.quot_thm_rty_qty quot_thm |> fst; | |
| 348 | val thy = Proof_Context.theory_of lthy; | |
| 349 | in | |
| 350 | if Sign.typ_instance thy (rty, T_of_bnf bnf_F) then | |
| 351 | (case pcr_info of | |
| 352 |             SOME {pcrel_def, ...} =>
 | |
| 353 | mk_transfer_thms' bnf_F bnf_G name consts Tss crel_def pcrel_def defs lthy | |
| 354 | | _ => maybe_warn_transfer pcr_why lthy) | |
| 355 | else maybe_warn_transfer (wrong_quotient rty) lthy | |
| 356 | end) | |
| 357 | | _ => maybe_warn_transfer no_quotient lthy | |
| 358 | end; | |
| 359 | ||
| 360 | ||
| 361 | (** typedef_bnf **) | |
| 362 | ||
| 363 | fun mk_typedef_transfer_tacs bnf_F bnf_G thms old_defs | |
| 364 | map_raw rel_raw pred_raw sets_raw = let | |
| 365 | ||
| 366 | val live = live_of_bnf bnf_G; | |
| 367 |     val Abs_G_inverse = @{thm type_definition.Abs_inverse} OF [#typedef thms];
 | |
| 368 |     val Rep_G = @{thm type_definition.Rep} OF [#typedef thms];
 | |
| 369 | ||
| 370 | fun common_tac addefs tac = (fn _ => fn ctxt => | |
| 371 | HEADGOAL (EVERY' [SELECT_GOAL (unfold_thms_tac ctxt addefs), | |
| 372 | REPEAT_DETERM o rtac ctxt rel_funI, | |
| 373 |         SELECT_GOAL (unfold_thms_tac ctxt @{thms o_apply}),
 | |
| 374 |         REPEAT_DETERM o eresolve_tac ctxt @{thms relcomppE exE conjE},
 | |
| 375 | hyp_subst_tac ctxt]) THEN tac ctxt) | |
| 376 | ||
| 377 | fun map_tac ctxt = (HEADGOAL o EVERY') | |
| 378 |       [rtac ctxt @{thm relcomppI},
 | |
| 379 | etac ctxt (mk_rel_funDN (live+1) (map_transfer_of_bnf bnf_F)), | |
| 380 | REPEAT_DETERM_N live o assume_tac ctxt, | |
| 381 | SELECT_GOAL (unfold_thms_tac ctxt [Abs_G_inverse OF [#map_closed thms] OF [Rep_G]]), | |
| 382 | REPEAT_DETERM o rtac ctxt refl]; | |
| 383 | val map_tac = common_tac [#map old_defs] map_tac; | |
| 384 | ||
| 385 | fun rel_tac ctxt = | |
| 386 | HEADGOAL (etac ctxt (mk_rel_funDN (live+2) (rel_transfer_of_bnf bnf_F)) THEN' | |
| 387 | REPEAT_DETERM_N (live+1) o assume_tac ctxt); | |
| 388 |     val rel_tac = common_tac (#rel old_defs :: @{thms vimage2p_def}) rel_tac;
 | |
| 389 | ||
| 390 | fun pred_tac ctxt = | |
| 391 | HEADGOAL (etac ctxt (mk_rel_funDN (live+1) (pred_transfer_of_bnf bnf_F)) THEN' | |
| 392 | REPEAT_DETERM_N live o (assume_tac ctxt)); | |
| 393 | val pred_tac = common_tac [#pred old_defs] pred_tac; | |
| 394 | ||
| 395 | fun set_tac set_transfer_thm ctxt = | |
| 396 | HEADGOAL (etac ctxt (rel_funD OF [set_transfer_thm])); | |
| 397 | fun mk_set_tac set_def set_transfer = common_tac [set_def] (set_tac set_transfer); | |
| 398 | val set_tacs = map2 mk_set_tac (#sets old_defs) (set_transfer_of_bnf bnf_F); | |
| 399 | ||
| 400 |   in {map={raw=map_raw,tac=map_tac},rel={raw=rel_raw,tac=rel_tac},
 | |
| 401 |       sets={raws=sets_raw,tacs=set_tacs},pred=SOME{raw=pred_raw,tac=pred_tac}} end;
 | |
| 60918 | 402 | |
| 62324 | 403 | fun typedef_bnf thm wits specs map_b rel_b pred_b opts lthy = | 
| 60918 | 404 | let | 
| 61073 | 405 | val plugins = | 
| 406 | get_first (fn Plugins_Option f => SOME (f lthy) | _ => NONE) (rev opts) | |
| 71262 | 407 | |> the_default Plugin_Name.default_filter; | 
| 60918 | 408 | |
| 409 | (* extract Rep Abs F RepT AbsT *) | |
| 71262 | 410 | val (Rep_G, Abs_G, F) = strip3 thm; | 
| 61073 | 411 | val typ_Abs_G = dest_funT (fastype_of Abs_G); | 
| 60918 | 412 | val RepT = fst typ_Abs_G; (* F *) | 
| 413 | val AbsT = snd typ_Abs_G; (* G *) | |
| 80634 
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
 wenzelm parents: 
79564diff
changeset | 414 | val AbsT_name = dest_Type_name AbsT; | 
| 
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
 wenzelm parents: 
79564diff
changeset | 415 | val tvs = AbsT |> dest_Type_args |> map (fst o dest_TVar); | 
| 60918 | 416 | val alpha0s = map (TFree o snd) specs; | 
| 61067 | 417 | |
| 62690 | 418 | val _ = length tvs = length alpha0s orelse | 
| 419 |       error ("Expected " ^ string_of_int (length tvs) ^ " type argument(s) to " ^ quote AbsT_name);
 | |
| 420 | ||
| 60918 | 421 | (* instantiate the new type variables newtvs to oldtvs *) | 
| 422 | val subst = subst_TVars (tvs ~~ alpha0s); | |
| 423 | val typ_subst = typ_subst_TVars (tvs ~~ alpha0s); | |
| 424 | ||
| 425 | val Rep_G = subst Rep_G; | |
| 426 | val Abs_G = subst Abs_G; | |
| 427 | val F = subst F; | |
| 428 | val RepT = typ_subst RepT; | |
| 429 | val AbsT = typ_subst AbsT; | |
| 430 | ||
| 61073 | 431 | fun flatten_tyargs Ass = | 
| 432 | map dest_TFree alpha0s | |
| 71262 | 433 | |> filter (fn T => exists (fn Ts => member op= Ts T) Ass); | 
| 60918 | 434 | |
| 435 | val Ds0 = filter (is_none o fst) specs |> map snd; | |
| 436 | ||
| 437 | (* get the bnf for RepT *) | |
| 71262 | 438 | val ((bnf_F, (deads, alphas)), ((_, unfolds), lthy)) = | 
| 62621 | 439 | bnf_of_typ true Dont_Inline (Binding.qualify true AbsT_name) flatten_tyargs [] | 
| 60918 | 440 | Ds0 RepT ((empty_comp_cache, empty_unfolds), lthy); | 
| 441 | ||
| 61073 | 442 | val set_bs = | 
| 443 | map (fn T => find_index (fn U => T = U) alpha0s) alphas | |
| 60918 | 444 | |> map (the_default Binding.empty o fst o nth specs); | 
| 445 | ||
| 61067 | 446 | val _ = (case alphas of [] => error "No live variables" | _ => ()); | 
| 60918 | 447 | |
| 71261 | 448 | val defs = #map_unfolds unfolds @ flat (#set_unfoldss unfolds) @ #rel_unfolds unfolds @ #pred_unfolds unfolds; | 
| 60918 | 449 | |
| 450 | (* number of live variables *) | |
| 71262 | 451 | val live = length alphas; | 
| 60918 | 452 | |
| 453 | (* state the three required properties *) | |
| 454 | val sorts = map Type.sort_of_atyp alphas; | |
| 79564 
33b10cd883ae
made lift_bnf more robust for abstract types with 'phantom' type variables
 traytel parents: 
75624diff
changeset | 455 | val names_lthy = fold Variable.declare_typ (alphas @ Library.union (op =) deads (map TFree Ds0)) lthy; | 
| 71262 | 456 | val (((alphas', betas), betas'), names_lthy) = names_lthy | 
| 457 | |> mk_TFrees' sorts | |
| 458 | ||>> mk_TFrees' sorts | |
| 459 | ||>> mk_TFrees' sorts; | |
| 60918 | 460 | |
| 71262 | 461 | val map_F = mk_map_of_bnf deads alphas betas bnf_F; | 
| 60918 | 462 | |
| 71262 | 463 | val (typ_fs, typ_aF) = fastype_of map_F |> strip_typeN live ||> domain_type; | 
| 60918 | 464 | val typ_pairs = map HOLogic.mk_prodT (alphas ~~ alphas'); | 
| 465 | val typ_subst_pair = typ_subst_atomic (alphas ~~ typ_pairs); | |
| 466 | val typ_pair = typ_subst_pair RepT; | |
| 467 | val subst_b = subst_atomic_types (alphas ~~ betas); | |
| 468 | val subst_a' = subst_atomic_types (alphas ~~ alphas'); | |
| 469 | val subst_pair = subst_atomic_types (alphas ~~ typ_pairs); | |
| 470 | val aF_set = F; | |
| 471 | val aF_set' = subst_a' F; | |
| 472 | val pairF_set = subst_pair F; | |
| 71262 | 473 | val bF_set = subst_b F; | 
| 474 | val map_F_fst = mk_map_of_bnf deads typ_pairs alphas bnf_F; | |
| 475 | val map_F_snd = mk_map_of_bnf deads typ_pairs alphas' bnf_F | |
| 476 | val sets_F_pairs = mk_sets_of_bnf (replicate live deads) (replicate live typ_pairs) bnf_F | |
| 60918 | 477 | val wits_F = mk_wits_of_bnf | 
| 71262 | 478 | (replicate (nwits_of_bnf bnf_F) deads) | 
| 479 | (replicate (nwits_of_bnf bnf_F) alphas) bnf_F; | |
| 60918 | 480 | |
| 481 |     (* val map_closed_F = @{term "\<And>f x. x \<in> F \<Longrightarrow> map_F f x \<in> F"}; *)
 | |
| 482 | val (var_fs, names_lthy) = mk_Frees "f" typ_fs names_lthy; | |
| 483 | val (var_x, names_lthy) = mk_Frees "x" [typ_aF] names_lthy |>> the_single; | |
| 61073 | 484 | val mem_x = HOLogic.mk_Trueprop (HOLogic.mk_mem (var_x, aF_set)); | 
| 60918 | 485 | val map_f = list_comb (map_F, var_fs); | 
| 61073 | 486 | val mem_map = HOLogic.mk_Trueprop (HOLogic.mk_mem (map_f $ var_x, bF_set)); | 
| 60918 | 487 | val imp_map = Logic.mk_implies (mem_x, mem_map); | 
| 61073 | 488 | val map_closed_F = fold_rev Logic.all var_fs (Logic.all var_x imp_map); | 
| 60918 | 489 | |
| 71262 | 490 | (* val zip_closed_F = | 
| 491 |       @{term "\<And>z. \<lbrakk>map_F fst z \<in> F; map_F snd z \<in> F\<rbrakk> \<Longrightarrow>
 | |
| 492 | \<exists>z' \<in> F. set_F z' \<subseteq> set_F z \<and> map_F fst z' = map_F fst z \<and> map_F snd z' = map_F snd z"}; *) | |
| 493 | val (var_z, names_lthy) = mk_Free "z" typ_pair names_lthy; | |
| 494 | val (var_z', names_lthy) = mk_Free "z'" typ_pair names_lthy; | |
| 60918 | 495 | val (pairs, names_lthy) = mk_Frees "tmp" typ_pairs names_lthy; | 
| 71262 | 496 | |
| 497 | fun mk_map mfs f z = Term.list_comb (mfs, map (fst o Term.strip_comb o f) pairs) $ z; | |
| 498 | fun mk_set var = map (fn t => t $ var) sets_F_pairs; | |
| 499 | ||
| 500 | val (map_fst', map_fst) = apply2 (mk_map map_F_fst HOLogic.mk_fst) (var_z', var_z); | |
| 501 | val (map_snd', map_snd) = apply2 (mk_map map_F_snd HOLogic.mk_snd) (var_z', var_z); | |
| 61073 | 502 | val mem_map_fst = HOLogic.mk_Trueprop (HOLogic.mk_mem (map_fst, aF_set)); | 
| 503 | val mem_map_snd = HOLogic.mk_Trueprop (HOLogic.mk_mem (map_snd, aF_set')); | |
| 71262 | 504 | val ex_conj = foldr1 HOLogic.mk_conj (map2 mk_leq (mk_set var_z') (mk_set var_z) @ | 
| 505 | [HOLogic.mk_eq (map_fst', map_fst), HOLogic.mk_eq (map_snd', map_snd)]); | |
| 506 | val zip_concl = HOLogic.mk_Trueprop (mk_Bex pairF_set (absfree (dest_Free var_z') ex_conj)); | |
| 507 | val zip_closed_F = Logic.all var_z (Logic.list_implies ([mem_map_fst, mem_map_snd], zip_concl)); | |
| 60918 | 508 | |
| 509 |     (* val wit_closed_F = @{term "wit_F a \<in> F"}; *)
 | |
| 510 | val (var_as, names_lthy) = mk_Frees "a" alphas names_lthy; | |
| 511 | val (var_bs, _) = mk_Frees "a" alphas names_lthy; | |
| 71262 | 512 | val sets = mk_sets_of_bnf (replicate live deads) (replicate live alphas) bnf_F; | 
| 513 | val (Iwits, wit_goals) = | |
| 514 | prepare_wits false RepT wits opts alphas wits_F var_as var_bs sets lthy; | |
| 60918 | 515 | val wit_closed_Fs = | 
| 61067 | 516 | Iwits |> map (fn (I, wit_F) => | 
| 60918 | 517 | let | 
| 518 | val vars = map (nth var_as) I; | |
| 519 | val wit_a = list_comb (wit_F, vars); | |
| 61073 | 520 | in fold_rev Logic.all vars (HOLogic.mk_Trueprop (HOLogic.mk_mem (wit_a, aF_set))) end); | 
| 60918 | 521 | |
| 522 | val goals = [map_closed_F, zip_closed_F] @ wit_closed_Fs @ | |
| 71262 | 523 | (case wits of NONE => [] | _ => wit_goals); | 
| 60918 | 524 | |
| 525 | fun after_qed ([map_closed_thm] :: [zip_closed_thm] :: wit_thmss) lthy = | |
| 61067 | 526 | let | 
| 527 | val (wit_closed_thms, wit_thms) = | |
| 528 | (case wits of | |
| 71262 | 529 | NONE => (map the_single wit_thmss, wit_thms_of_bnf bnf_F) | 
| 61067 | 530 | | _ => chop (length wit_closed_Fs) (map the_single wit_thmss)) | 
| 60918 | 531 | |
| 61067 | 532 | (* construct map set bd rel wit *) | 
| 533 |             (* val map_G = @{term "\<lambda>f. Abs_G o map_F f o Rep_G"}; *)
 | |
| 534 | val Abs_Gb = subst_b Abs_G; | |
| 71262 | 535 | val map_G = fold_rev lambda var_fs | 
| 61073 | 536 | (HOLogic.mk_comp (HOLogic.mk_comp (Abs_Gb, map_f), Rep_G)); | 
| 71262 | 537 | val map_raw = fold_rev lambda var_fs map_f; | 
| 60918 | 538 | |
| 61067 | 539 |             (* val sets_G = [@{term "set_F o Rep_G"}]; *)
 | 
| 71262 | 540 | val sets_F = mk_sets_of_bnf (replicate live deads) (replicate live alphas) bnf_F; | 
| 61067 | 541 | val sets_G = map (fn set_F => HOLogic.mk_comp (set_F, Rep_G)) sets_F; | 
| 60918 | 542 | |
| 61067 | 543 |             (* val bd_G = @{term "bd_F"}; *)
 | 
| 71262 | 544 | val bd_F = mk_bd_of_bnf deads alphas bnf_F; | 
| 61067 | 545 | val bd_G = bd_F; | 
| 546 | ||
| 547 |             (* val rel_G = @{term "\<lambda>R. BNF_Def.vimage2p Rep_G Rep_G (rel_F R)"}; *)
 | |
| 71262 | 548 | val rel_F = mk_rel_of_bnf deads alphas betas bnf_F; | 
| 549 | val (typ_Rs, _) = strip_typeN live (fastype_of rel_F); | |
| 60918 | 550 | |
| 61067 | 551 | val (var_Rs, names_lthy) = mk_Frees "R" typ_Rs lthy; | 
| 552 | val Rep_Gb = subst_b Rep_G; | |
| 553 | val rel_G = fold_rev absfree (map dest_Free var_Rs) | |
| 554 | (mk_vimage2p Rep_G Rep_Gb $ list_comb (rel_F, var_Rs)); | |
| 71262 | 555 | val rel_raw = fold_rev absfree (map dest_Free var_Rs) (list_comb (rel_F, var_Rs)); | 
| 60918 | 556 | |
| 62324 | 557 |             (* val pred_G = @{term "\<lambda>P. pred_F P o Rep_G"}; *)
 | 
| 71262 | 558 | val pred_F = mk_pred_of_bnf deads alphas bnf_F; | 
| 559 | val (typ_Ps, _) = strip_typeN live (fastype_of pred_F); | |
| 62324 | 560 | |
| 561 | val (var_Ps, names_lthy) = mk_Frees "P" typ_Ps names_lthy; | |
| 562 | val pred_G = fold_rev absfree (map dest_Free var_Ps) | |
| 563 | (HOLogic.mk_comp (list_comb (pred_F, var_Ps), Rep_G)); | |
| 71262 | 564 | val pred_raw = fold_rev absfree (map dest_Free var_Ps) (list_comb (pred_F, var_Ps)); | 
| 62324 | 565 | |
| 61067 | 566 |             (* val wits_G = [@{term "Abs_G o wit_F"}]; *)
 | 
| 567 | val (var_as, _) = mk_Frees "a" alphas names_lthy; | |
| 568 | val wits_G = | |
| 569 | map (fn (I, wit_F) => | |
| 570 | let | |
| 571 | val vs = map (nth var_as) I; | |
| 572 | in fold_rev absfree (map dest_Free vs) (Abs_G $ (list_comb (wit_F, vs))) end) | |
| 573 | Iwits; | |
| 60918 | 574 | |
| 61067 | 575 | (* tactics *) | 
| 576 |             val Rep_thm = thm RS @{thm type_definition.Rep};
 | |
| 577 |             val Abs_inverse_thm = thm RS @{thm type_definition.Abs_inverse};
 | |
| 578 |             val Abs_inject_thm = thm RS @{thm type_definition.Abs_inject};
 | |
| 579 |             val Rep_cases_thm = thm RS @{thm type_definition.Rep_cases};
 | |
| 580 |             val Rep_inverse_thm = thm RS @{thm type_definition.Rep_inverse};
 | |
| 60918 | 581 | |
| 61067 | 582 | fun map_id0_tac ctxt = | 
| 583 | HEADGOAL (EVERY' [rtac ctxt ext, | |
| 71262 | 584 | SELECT_GOAL (unfold_thms_tac ctxt [map_id0_of_bnf bnf_F, id_apply, o_apply, | 
| 61067 | 585 | Rep_inverse_thm]), | 
| 586 | rtac ctxt refl]); | |
| 60918 | 587 | |
| 61067 | 588 | fun map_comp0_tac ctxt = | 
| 60918 | 589 | HEADGOAL (EVERY' [rtac ctxt ext, | 
| 71262 | 590 | SELECT_GOAL (unfold_thms_tac ctxt [map_comp0_of_bnf bnf_F, o_apply, | 
| 60918 | 591 | Rep_thm RS (map_closed_thm RS Abs_inverse_thm)]), | 
| 61067 | 592 | rtac ctxt refl]); | 
| 60918 | 593 | |
| 61067 | 594 | fun map_cong0_tac ctxt = | 
| 595 | HEADGOAL (EVERY' ([SELECT_GOAL (unfold_thms_tac ctxt [o_apply]), | |
| 596 | rtac ctxt (([Rep_thm RS map_closed_thm, Rep_thm RS map_closed_thm] MRS | |
| 597 | Abs_inject_thm) RS iffD2), | |
| 71262 | 598 | rtac ctxt (map_cong0_of_bnf bnf_F)] @ replicate live (Goal.assume_rule_tac ctxt))); | 
| 60918 | 599 | |
| 61067 | 600 | val set_map0s_tac = | 
| 601 | map (fn set_map => fn ctxt => | |
| 602 | HEADGOAL (EVERY' [rtac ctxt ext, | |
| 603 | SELECT_GOAL (unfold_thms_tac ctxt [set_map, o_apply, | |
| 604 | Rep_thm RS (map_closed_thm RS Abs_inverse_thm)]), | |
| 605 | rtac ctxt refl])) | |
| 71262 | 606 | (set_map_of_bnf bnf_F); | 
| 61067 | 607 | |
| 71262 | 608 | fun card_order_bd_tac ctxt = HEADGOAL (rtac ctxt (bd_card_order_of_bnf bnf_F)); | 
| 60918 | 609 | |
| 71262 | 610 | fun cinfinite_bd_tac ctxt = HEADGOAL (rtac ctxt (bd_cinfinite_of_bnf bnf_F)); | 
| 61067 | 611 | |
| 75624 | 612 | fun regularCard_bd_tac ctxt = HEADGOAL (rtac ctxt (bd_regularCard_of_bnf bnf_F)); | 
| 613 | ||
| 61067 | 614 | val set_bds_tac = | 
| 615 | map (fn set_bd => fn ctxt => | |
| 616 | HEADGOAL (EVERY' [SELECT_GOAL (unfold_thms_tac ctxt [o_apply]), rtac ctxt set_bd])) | |
| 71262 | 617 | (set_bd_of_bnf bnf_F); | 
| 61067 | 618 | |
| 619 | fun le_rel_OO_tac ctxt = | |
| 620 |               HEADGOAL (EVERY' [rtac ctxt @{thm vimage2p_relcompp_mono},
 | |
| 71262 | 621 |                 rtac ctxt ((rel_OO_of_bnf bnf_F RS sym) RS @{thm ord_eq_le_trans}),
 | 
| 61067 | 622 |                 rtac ctxt @{thm order_refl}]);
 | 
| 60918 | 623 | |
| 61067 | 624 | fun rel_OO_Grp_tac ctxt = | 
| 625 | HEADGOAL (EVERY' ([SELECT_GOAL (REPEAT_DETERM (HEADGOAL (rtac ctxt ext))), | |
| 626 |                 SELECT_GOAL (unfold_thms_tac ctxt [@{thm OO_Grp_alt}, mem_Collect_eq,
 | |
| 71262 | 627 |                   o_apply, @{thm vimage2p_def}, in_rel_of_bnf bnf_F, Bex_def, mem_Collect_eq]),
 | 
| 61067 | 628 | rtac ctxt iffI, | 
| 71262 | 629 | SELECT_GOAL (REPEAT_DETERM (HEADGOAL (eresolve_tac ctxt [exE, conjE]))), | 
| 630 | forward_tac ctxt | |
| 631 |                   [zip_closed_thm OF (replicate 2 (Rep_thm RSN (2, @{thm ssubst_mem})))],
 | |
| 61067 | 632 | assume_tac ctxt, | 
| 71262 | 633 | SELECT_GOAL (REPEAT_DETERM (HEADGOAL (eresolve_tac ctxt [bexE, conjE]))), | 
| 634 | etac ctxt Rep_cases_thm, | |
| 61067 | 635 | hyp_subst_tac ctxt, | 
| 636 | SELECT_GOAL (REPEAT_DETERM (HEADGOAL (rtac ctxt exI))), | |
| 637 | rtac ctxt conjI] @ | |
| 71262 | 638 | replicate live | 
| 639 |                   (EVERY' [TRY o rtac ctxt conjI, etac ctxt @{thm order_trans}, assume_tac ctxt]) @
 | |
| 640 | [SELECT_GOAL (REPEAT_DETERM (HEADGOAL (rtac ctxt conjI))), | |
| 641 | REPEAT_DETERM_N 2 o EVERY' | |
| 642 | [rtac ctxt (trans OF [iffD2 OF [Abs_inject_thm OF | |
| 643 | [map_closed_thm OF [Rep_thm], Rep_thm]], Rep_inverse_thm]), | |
| 644 | etac ctxt trans, assume_tac ctxt], | |
| 645 | SELECT_GOAL (REPEAT_DETERM (HEADGOAL (eresolve_tac ctxt [exE, conjE]))), | |
| 61067 | 646 | rtac ctxt exI, | 
| 647 | rtac ctxt conjI] @ | |
| 71262 | 648 | replicate (live-1) (rtac ctxt conjI THEN' assume_tac ctxt) @ | 
| 61067 | 649 | [assume_tac ctxt, | 
| 650 | rtac ctxt conjI, | |
| 651 | REPEAT_DETERM_N 2 o EVERY' | |
| 652 | [rtac ctxt (iffD1 OF [Abs_inject_thm OF [map_closed_thm OF [Rep_thm], Rep_thm]]), | |
| 653 | etac ctxt (Rep_inverse_thm RS sym RSN (2, trans))]])); | |
| 60918 | 654 | |
| 62324 | 655 | fun pred_set_tac ctxt = | 
| 656 | HEADGOAL (EVERY' | |
| 71262 | 657 |                 [rtac ctxt (pred_set_of_bnf bnf_F RS @{thm arg_cong[of _ _ "\<lambda>f. f \<circ> _"]} RS trans),
 | 
| 62324 | 658 |                 SELECT_GOAL (unfold_thms_tac ctxt (@{thms Ball_comp_iff conj_comp_iff})),
 | 
| 659 | rtac ctxt refl]); | |
| 660 | ||
| 61067 | 661 | fun wit_tac ctxt = | 
| 662 | HEADGOAL (EVERY' | |
| 663 | (map (fn thm => (EVERY' | |
| 664 | [SELECT_GOAL (unfold_thms_tac ctxt (o_apply :: | |
| 665 | (wit_closed_thms RL [Abs_inverse_thm]))), | |
| 666 | dtac ctxt thm, assume_tac ctxt])) | |
| 667 | wit_thms)); | |
| 60918 | 668 | |
| 61067 | 669 | val tactics = [map_id0_tac, map_comp0_tac, map_cong0_tac] @ set_map0s_tac @ | 
| 75624 | 670 | [card_order_bd_tac, cinfinite_bd_tac, regularCard_bd_tac] @ set_bds_tac @ | 
| 62324 | 671 | [le_rel_OO_tac, rel_OO_Grp_tac, pred_set_tac]; | 
| 60918 | 672 | |
| 71262 | 673 | val (bnf_G, lthy) = bnf_def Dont_Inline (user_policy Note_Some) true I | 
| 62324 | 674 | tactics wit_tac NONE map_b rel_b pred_b set_bs | 
| 675 | (((((((Binding.empty, AbsT), map_G), sets_G), bd_G), wits_G), SOME rel_G), SOME pred_G) | |
| 61067 | 676 | lthy; | 
| 60928 
141a1d485259
unfold intermediate definitions (stemming from composition) in lifted bnf operations
 traytel parents: 
60918diff
changeset | 677 | |
| 71262 | 678 | val old_defs = | 
| 679 |               {sets = set_defs_of_bnf bnf_G, map = map_def_of_bnf bnf_G, rel = rel_def_of_bnf bnf_G,
 | |
| 680 | pred = pred_def_of_bnf bnf_G}; | |
| 681 | ||
| 682 | val unfold_morphism = Morphism.thm_morphism "BNF" (unfold_thms lthy defs); | |
| 683 | val (bnf_G, lthy) = morph_bnf_defs unfold_morphism bnf_G | |
| 66272 
c6714a9562ae
store the unfolded definitions of the lifted bnf constants under "_def" name
 traytel parents: 
63023diff
changeset | 684 | |> (fn bnf => note_bnf_defs bnf lthy); | 
| 71262 | 685 | |
| 686 | val quiet = exists (fn No_Warn_Transfer => true | _ => false) opts; | |
| 687 | ||
| 688 | val transfer_consts = mk_typedef_transfer_tacs bnf_F bnf_G | |
| 689 |               {map_closed=map_closed_thm,typedef=thm} old_defs map_raw rel_raw pred_raw sets_F;
 | |
| 61067 | 690 | in | 
| 71262 | 691 | lthy |> BNF_Def.register_bnf plugins AbsT_name bnf_G |> | 
| 71494 | 692 | mk_transfer_thms quiet bnf_F bnf_G AbsT_name transfer_consts Typedef | 
| 71262 | 693 |               {abs=typ_subst_atomic (alphas ~~ alphas') AbsT, rep=RepT, Ds0=map TFree Ds0,
 | 
| 694 | deads = deads, alphas=alphas, betas=alphas', gammas=betas, deltas=betas'} defs | |
| 61067 | 695 | end | 
| 696 | | after_qed _ _ = raise Match; | |
| 60918 | 697 | in | 
| 698 | (goals, after_qed, defs, lthy) | |
| 699 | end; | |
| 700 | ||
| 61067 | 701 | |
| 71262 | 702 | (** quotient_bnf **) | 
| 703 | ||
| 704 | fun mk_quotient_transfer_tacs bnf_F Tss live qthms thms set_F'_thmss old_defs | |
| 705 | inst_REL_pos_distrI map_raw rel_raw sets_raw = let | |
| 706 | ||
| 707 | fun common_tac ctxt addefs = unfold_thms_tac ctxt (#REL qthms :: addefs) THEN | |
| 708 | (REPEAT_DETERM o HEADGOAL) (rtac ctxt rel_funI); | |
| 709 | ||
| 710 | (* quotient.map_transfer tactic *) | |
| 711 | val map_F_transfer = map_transfer_of_bnf bnf_F |> mk_rel_funDN (live+1); | |
| 712 | fun map_transfer_q _ ctxt = | |
| 713 |       common_tac ctxt (#map old_defs :: @{thms o_def}) THEN
 | |
| 714 |       (HEADGOAL o EVERY') [REPEAT_DETERM o eresolve_tac ctxt @{thms relcomppE},
 | |
| 715 |         rtac ctxt @{thm relcomppI[rotated]}, hyp_subst_tac ctxt THEN' EVERY'
 | |
| 716 | (map (rtac ctxt) [#rel_abs qthms, #map_F_rsp thms, (#rep_abs_rsp qthms), (#reflp qthms)]), | |
| 717 | hyp_subst_tac ctxt, rtac ctxt map_F_transfer, REPEAT_DETERM_N (live+1) o assume_tac ctxt]; | |
| 718 | ||
| 719 | (* quotient.rel_transfer tactic *) | |
| 720 | val rel_F_maps = rel_map_of_bnf bnf_F; | |
| 721 |     val rel_F_map_iffD2s = map (fn thm => thm RS @{thm iffD2}) rel_F_maps;
 | |
| 722 | fun inst_REL_pos_distrI_order_refls vs aTs bTs ctxt = inst_REL_pos_distrI live vs aTs bTs ctxt | |
| 723 |       OF (replicate (live+1) asm_rl @ replicate live @{thm order_refl});
 | |
| 724 |     fun rel_transfer_q {Qs, Rs} ctxt = EVERY
 | |
| 725 |       [common_tac ctxt [#rel old_defs, @{thm vimage2p_def}],
 | |
| 726 | HEADGOAL (rtac ctxt iffI), | |
| 727 | (REPEAT_DETERM o ALLGOALS) | |
| 728 |         (eresolve_tac ctxt @{thms exE conjE relcomppE} ORELSE' hyp_subst_tac ctxt),
 | |
| 729 | (HEADGOAL o EVERY') | |
| 730 |         [REPEAT_DETERM o dtac ctxt @{thm rel_fun_rel_OO1},
 | |
| 731 | rtac ctxt (inst_REL_pos_distrI 0 (map mk_conversep Qs) (#betas Tss) (#alphas Tss) ctxt), | |
| 732 |         rtac ctxt @{thm relcomppI},
 | |
| 733 | rtac ctxt (#symp qthms), | |
| 734 | rtac ctxt (#map_F_rsp thms), | |
| 735 | rtac ctxt (#rep_abs_rsp qthms), | |
| 736 | rtac ctxt (#reflp qthms), | |
| 737 |         rtac ctxt @{thm relcomppI},
 | |
| 738 |         rtac ctxt (rel_flip_of_bnf bnf_F RS @{thm iffD1}),
 | |
| 739 | rtac ctxt (nth rel_F_map_iffD2s 0), | |
| 740 | rtac ctxt (nth rel_F_map_iffD2s 1), | |
| 741 | etac ctxt (#rel_monoD_rotated thms)], | |
| 742 | (REPEAT_DETERM_N live o HEADGOAL o EVERY') | |
| 743 |         [rtac ctxt @{thm predicate2I},
 | |
| 744 |         rtac ctxt @{thm conversepI},
 | |
| 745 |         rtac ctxt @{thm Basic_BNFs.rel_sum_simps(4)[THEN iffD2]},
 | |
| 746 |         etac ctxt @{thm conversepI}],
 | |
| 747 | (HEADGOAL o EVERY') | |
| 748 | [rtac ctxt (inst_REL_pos_distrI_order_refls Rs (#gammas Tss) (#deltas Tss) ctxt), | |
| 749 |         (SELECT_GOAL o REPEAT_DETERM o HEADGOAL) (etac ctxt @{thm relcomppI}),
 | |
| 750 |         rtac ctxt @{thm relcomppI[rotated]},
 | |
| 751 | rtac ctxt (#map_F_rsp thms), | |
| 752 | rtac ctxt (#rep_abs_rsp qthms OF [#reflp qthms]), | |
| 753 |         SELECT_GOAL (unfold_thms_tac ctxt (@{thms rel_sum_simps} @ rel_F_maps)),
 | |
| 754 | assume_tac ctxt], | |
| 755 | (REPEAT_DETERM_N (2*live) o HEADGOAL) | |
| 756 |         (rtac ctxt @{thm rel_sum_eq2_nonempty} ORELSE' rtac ctxt @{thm rel_sum_eq3_nonempty}),
 | |
| 757 | (REPEAT_DETERM_N live) | |
| 758 |         (unfold_thms_tac ctxt @{thms sum.rel_compp[symmetric] eq_OO} THEN
 | |
| 759 |         HEADGOAL (etac ctxt @{thm sum.rel_mono[OF order_refl]})),
 | |
| 760 | (HEADGOAL o EVERY') | |
| 761 |         [(SELECT_GOAL o REPEAT_DETERM o HEADGOAL) (dtac ctxt @{thm rel_fun_rel_OO2}),
 | |
| 762 | rtac ctxt (inst_REL_pos_distrI 0 Qs (#alphas Tss) (#betas Tss) ctxt), | |
| 763 |         rtac ctxt @{thm relcomppI},
 | |
| 764 | rtac ctxt (#reflp qthms), | |
| 765 |         rtac ctxt @{thm relcomppI},
 | |
| 766 | rtac ctxt (nth rel_F_map_iffD2s 0), | |
| 767 | rtac ctxt (nth rel_F_map_iffD2s 1), | |
| 768 | etac ctxt (#rel_monoD_rotated thms)], | |
| 769 | (REPEAT_DETERM_N live o HEADGOAL o EVERY') | |
| 770 |         [rtac ctxt @{thm predicate2I}, rtac ctxt @{thm rel_sum.intros(2)}, assume_tac ctxt],
 | |
| 771 | (HEADGOAL o EVERY') | |
| 772 | [rtac ctxt | |
| 773 | (inst_REL_pos_distrI_order_refls (map mk_conversep Rs) (#deltas Tss) (#gammas Tss) ctxt), | |
| 774 |         rtac ctxt @{thm relcomppI},
 | |
| 775 | etac ctxt (rotate_prems 1 (#transp qthms)), | |
| 776 | rtac ctxt (#map_F_rsp thms), | |
| 777 | rtac ctxt (#rep_abs_rsp qthms OF [#reflp qthms]), | |
| 778 |         etac ctxt @{thm relcomppI},
 | |
| 779 |         rtac ctxt @{thm relcomppI},
 | |
| 780 | etac ctxt (#transp qthms), | |
| 781 | rtac ctxt (#symp qthms), | |
| 782 | rtac ctxt (#map_F_rsp thms), | |
| 783 | rtac ctxt (#rep_abs_rsp qthms), | |
| 784 | rtac ctxt (#reflp qthms), | |
| 785 |         rtac ctxt @{thm relcomppI[rotated]},
 | |
| 786 | rtac ctxt (#reflp qthms), | |
| 787 |         rtac ctxt (rel_flip_of_bnf bnf_F RS @{thm iffD1}),
 | |
| 788 | rtac ctxt (nth rel_F_map_iffD2s 0), | |
| 789 | rtac ctxt (nth rel_F_map_iffD2s 1), | |
| 790 | etac ctxt (#rel_monoD_rotated thms)], | |
| 791 | (REPEAT_DETERM_N live o HEADGOAL o EVERY') | |
| 792 |         [rtac ctxt @{thm predicate2I},
 | |
| 793 |         rtac ctxt @{thm conversepI},
 | |
| 794 |         rtac ctxt @{thm rel_sum.intros(2)},
 | |
| 795 |         etac ctxt @{thm conversepI}],
 | |
| 796 | (REPEAT_DETERM_N (2*live) o HEADGOAL) | |
| 797 |         (rtac ctxt @{thm rel_sum_eq2_nonempty} ORELSE' rtac ctxt @{thm rel_sum_eq3_nonempty}),
 | |
| 798 | (REPEAT_DETERM_N live o EVERY) | |
| 799 |         [unfold_thms_tac ctxt @{thms sum.rel_compp[symmetric] eq_OO},
 | |
| 800 |         HEADGOAL (etac ctxt @{thm sum.rel_mono[OF order_refl]})]];
 | |
| 801 | ||
| 802 | (* quotient.set_transfer tactics *) | |
| 803 | fun set_transfer_q set_G_def set_F'_thms _ ctxt = | |
| 804 | let val set_F'_rsp = mk_rel_funDN 1 (#set_F'_respect set_F'_thms) in | |
| 805 |         common_tac ctxt (set_G_def :: @{thms o_def}) THEN
 | |
| 806 | (HEADGOAL o EVERY') | |
| 807 |           [etac ctxt @{thm relcomppE}, hyp_subst_tac ctxt,
 | |
| 808 | SELECT_GOAL (unfold_thms_tac ctxt | |
| 809 |             [set_F'_rsp OF [#rep_abs qthms] OF [#reflp qthms], @{thm rel_set_def}]),
 | |
| 810 | dtac ctxt (#rel_F_rel_F' thms), rtac ctxt conjI] THEN | |
| 811 | (REPEAT_DETERM_N 2 o HEADGOAL o EVERY') | |
| 812 | [SELECT_GOAL (unfold_thms_tac ctxt [#rel_F'_set thms]), | |
| 813 | REPEAT_DETERM o eresolve_tac ctxt [exE, conjE], | |
| 814 | REPEAT_DETERM o dtac ctxt (mk_sym set_F'_rsp), | |
| 815 | SELECT_GOAL (unfold_thms_tac ctxt [#set_map_F' set_F'_thms]), | |
| 816 |           rtac ctxt ballI, dtac ctxt @{thm equalityD1[THEN subsetD]}, assume_tac ctxt,
 | |
| 817 |           SELECT_GOAL (unfold_thms_tac ctxt @{thms image_iff}),
 | |
| 818 | etac ctxt bexE, dtac ctxt set_mp, assume_tac ctxt, | |
| 819 |           SELECT_GOAL (unfold_thms_tac ctxt @{thms mem_Collect_eq case_prod_beta}),
 | |
| 820 |           rtac ctxt bexI, hyp_subst_tac ctxt, assume_tac ctxt, etac ctxt @{thm hypsubst},
 | |
| 821 |           etac ctxt @{thm imageI}, assume_tac ctxt]
 | |
| 822 | end; | |
| 823 | in | |
| 824 |      {map={raw=map_raw, tac=map_transfer_q},
 | |
| 825 |       rel={raw=rel_raw, tac=rel_transfer_q},
 | |
| 826 |       sets={raws=sets_raw,tacs=map2 set_transfer_q (#sets old_defs) set_F'_thmss},
 | |
| 827 | pred=NONE} | |
| 828 | end; | |
| 829 | ||
| 830 | ||
| 71494 | 831 | fun quotient_bnf (equiv_thm, quot_thm) wits specs map_b rel_b pred_b opts lthy = | 
| 71262 | 832 | let | 
| 833 | (* extract rep_G and abs_G *) | |
| 71393 
fce780f9c9c6
new examples of BNF lifting across quotients using a new theory of confluence,
 traytel parents: 
71354diff
changeset | 834 | val (equiv_rel, abs_G, rep_G) = strip3 quot_thm; | 
| 71262 | 835 |     val (repT, absT) = dest_funT (fastype_of abs_G); (* ("?'a F", "?'a G") *)
 | 
| 80634 
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
 wenzelm parents: 
79564diff
changeset | 836 | val absT_name = dest_Type_name absT; | 
| 71262 | 837 | |
| 80634 
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
 wenzelm parents: 
79564diff
changeset | 838 | val tvs = map (fst o dest_TVar) (dest_Type_args absT); | 
| 71262 | 839 | val _ = length tvs = length specs orelse | 
| 840 |       error ("Expected " ^ string_of_int (length tvs) ^
 | |
| 841 | " type argument" ^ (if (length tvs) = 1 then "" else "s") ^ " to " ^ quote absT_name); | |
| 842 | ||
| 843 | (* instantiate TVars *) | |
| 844 | val alpha0s = map (TFree o snd) specs; | |
| 845 | val typ_subst = typ_subst_TVars (tvs ~~ alpha0s); | |
| 846 | val (repT, absT) = apply2 typ_subst (repT, absT); | |
| 847 | ||
| 848 | (* get the bnf for RepT *) | |
| 849 | val Ds0 = filter (is_none o fst) specs |> map snd; | |
| 850 | ||
| 851 | fun flatten_tyargs Ass = | |
| 852 | map dest_TFree alpha0s |> filter (fn T => exists (fn Ts => member op= Ts T) Ass); | |
| 853 | ||
| 854 | val ((bnf_F, (deads, alphas)), ((_, unfolds), lthy)) = | |
| 855 | bnf_of_typ true Dont_Inline (Binding.qualify true absT_name) flatten_tyargs | |
| 856 | [] Ds0 repT ((empty_comp_cache, empty_unfolds), lthy); | |
| 857 | val live = length alphas; | |
| 858 | val _ = (if live = 0 then error "No live variables" else ()); | |
| 859 | ||
| 860 | val defs = #map_unfolds unfolds @ flat (#set_unfoldss unfolds) @ #rel_unfolds unfolds; | |
| 861 | val set_bs = | |
| 862 | map (fn T => find_index (fn U => T = U) alpha0s) alphas | |
| 863 | |> map (the_default Binding.empty o fst o nth specs); | |
| 864 | ||
| 865 | (* create and instantiate all the needed type variables *) | |
| 866 | val subst = subst_TVars (tvs ~~ alpha0s); | |
| 867 | val (abs_G, rep_G) = apply2 subst (abs_G, rep_G); | |
| 868 | ||
| 869 | val sorts = map Type.sort_of_atyp alphas; | |
| 870 | val (((betas, gammas), deltas), names_lthy) = fold Variable.declare_typ (alphas @ deads) lthy | |
| 871 | |> mk_TFrees' sorts | |
| 872 | ||>> mk_TFrees' sorts | |
| 873 | ||>> mk_TFrees' sorts; | |
| 874 | ||
| 875 | fun subst_Ts tm Ts = subst_atomic_types (alphas ~~ Ts) tm; | |
| 876 | val subst_b = subst_atomic_types (alphas ~~ betas); | |
| 877 | val subst_Maybe = subst_atomic_types o map (swap o `mk_MaybeT); | |
| 878 | val equiv_rel_a = subst equiv_rel; | |
| 879 | val map_F = mk_map_of_bnf deads alphas betas bnf_F; | |
| 880 | val rel_F_ab = mk_rel_of_bnf deads alphas betas bnf_F; | |
| 881 | val rel_F_bc = mk_rel_of_bnf deads betas gammas bnf_F; | |
| 882 | val rel_F_ac = mk_rel_of_bnf deads alphas gammas bnf_F; | |
| 883 | val rel_F_option = mk_rel_of_bnf deads (map mk_MaybeT alphas) (map mk_MaybeT betas) bnf_F; | |
| 884 | val sets_F = mk_sets_of_bnf (replicate live deads) (replicate live alphas) bnf_F; | |
| 885 | val wits_F = mk_wits_of_bnf | |
| 886 | (replicate (nwits_of_bnf bnf_F) deads) (replicate (nwits_of_bnf bnf_F) alphas) bnf_F; | |
| 887 | ||
| 888 | val (typ_fs, (typ_aF, typ_bF)) = strip_typeN live (fastype_of map_F) ||> dest_funT; | |
| 889 | val typ_MaybeF = typ_subst_atomic (alphas ~~ map mk_MaybeT alphas) typ_aF; | |
| 890 | val typ_a_sets = map HOLogic.mk_setT alphas; | |
| 891 | val typ_pairs = map HOLogic.mk_prodT (alphas ~~ betas); | |
| 892 | val typ_fs' = map (typ_subst_atomic (map (swap o `mk_MaybeT) betas)) typ_fs; | |
| 893 | ||
| 894 | (* create all the needed variables *) | |
| 895 | val ((((((((((((((((((((((var_Ps, var_Qs), var_Rs), var_x), var_x'), var_y), var_y'), var_mx), | |
| 896 | var_As), var_As'), var_Ss), var_Bs), var_as), var_as'), var_bs), var_bs'), var_R), var_fs), | |
| 897 | var_fs'), var_gs), var_gs'), var_z), var_ts) = names_lthy | |
| 898 | |> mk_Frees "Ps" (map2 mk_relT alphas betas) | |
| 899 | ||>> mk_Frees "Qs" (map2 mk_relT betas gammas) | |
| 900 | ||>> mk_Frees "Rs" (map2 mk_relT alphas gammas) | |
| 901 | ||>> mk_Free "x" typ_aF | |
| 902 | ||>> mk_Free "x'" typ_aF | |
| 903 | ||>> mk_Free "y" typ_bF | |
| 904 | ||>> mk_Free "y'" (typ_subst_atomic (alphas ~~ gammas) typ_aF) | |
| 905 | ||>> mk_Free "mx" typ_MaybeF | |
| 906 | ||>> mk_Frees "As" typ_a_sets | |
| 907 | ||>> mk_Frees "As'" typ_a_sets | |
| 908 | ||>> mk_Frees "Ss" (map HOLogic.mk_setT typ_a_sets) | |
| 909 | ||>> mk_Frees "Bs" (map HOLogic.mk_setT betas) | |
| 910 | ||>> mk_Frees "as" alphas | |
| 911 | ||>> mk_Frees "as'" alphas | |
| 912 | ||>> mk_Frees "bs" betas | |
| 913 | ||>> mk_Frees "bs'" betas | |
| 914 | ||>> mk_Free "R" (typ_aF --> typ_bF --> HOLogic.boolT) | |
| 915 | ||>> mk_Frees "fs" typ_fs | |
| 916 | ||>> mk_Frees "fs'" typ_fs' | |
| 917 | ||>> mk_Frees "gs" typ_fs | |
| 918 | ||>> mk_Frees "gs'" typ_fs' | |
| 919 | ||>> mk_Free "z" (typ_subst_atomic (alphas ~~ typ_pairs) typ_aF) | |
| 920 | ||>> mk_Frees "ts" typ_pairs | |
| 921 | |> fst; | |
| 922 | ||
| 923 | (* create local definitions `b = tm` with n arguments *) | |
| 80636 
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
 wenzelm parents: 
80634diff
changeset | 924 | fun suffix tm s = Long_Name.base_name (dest_Const_name tm) ^ s; | 
| 71262 | 925 | fun define lthy b n tm = | 
| 926 | let | |
| 927 | val b = Binding.qualify true absT_name (Binding.qualified_name b); | |
| 928 | val ((tm, (_, def)), (lthy, lthy_old)) = lthy | |
| 72450 | 929 | |> (snd o Local_Theory.begin_nested) | 
| 71262 | 930 | |> Local_Theory.define_internal (((Binding.concealed b, NoSyn), (Binding.empty_atts, tm))) | 
| 72450 | 931 | ||> `Local_Theory.end_nested; | 
| 71262 | 932 | val phi = Proof_Context.export_morphism lthy_old lthy; | 
| 933 | val tm = Term.subst_atomic_types | |
| 934 | (map (`(Morphism.typ phi)) (alphas @ betas @ gammas @ map TFree Ds0)) | |
| 935 | (Morphism.term phi tm); | |
| 936 | val def = mk_unabs_def n (HOLogic.mk_obj_eq (Morphism.thm phi def)); | |
| 937 |       in ({def=def, tm=tm}, lthy) end;
 | |
| 938 | ||
| 939 | (* internally use REL, not the user-provided definition *) | |
| 940 | val (REL, lthy) = define lthy "REL" 0 equiv_rel_a; | |
| 941 | val REL_def = sym RS eq_reflection OF [#def REL]; | |
| 942 | fun REL_rewr_all ctxt thm = Conv.fconv_rule | |
| 943 | (Conv.top_conv (fn _ => Conv.try_conv (Conv.rewr_conv REL_def)) ctxt) thm; | |
| 944 | ||
| 945 | val equiv_rel_a' = equiv_rel_a; | |
| 946 | val equiv_rel_a = #tm REL; | |
| 947 | val (equiv_rel_b, equiv_rel_c) = apply2 (subst_Ts equiv_rel_a) (betas, gammas); | |
| 948 | ||
| 949 |     (* rel_pos_distr: @{term "\<And>A B.
 | |
| 950 | A \<circ>\<circ> B \<noteq> bot \<Longrightarrow> REL \<circ>\<circ> rel_F A \<circ>\<circ> REL \<circ>\<circ> rel_F B \<circ>\<circ> REL \<le> REL \<circ>\<circ> rel_F (A \<circ>\<circ> B) \<circ>\<circ> REL"} *) | |
| 951 | fun compp_not_bot comp aT cT = let | |
| 952 | val T = mk_relT aT cT; | |
| 953 | val mk_eq = HOLogic.eq_const T; | |
| 954 | in HOLogic.mk_not (mk_eq $ comp $ bot_const T) end; | |
| 955 | val ab_comps = map2 mk_relcompp var_Ps var_Qs; | |
| 956 |     val ne_comps = (@{map 3} compp_not_bot ab_comps alphas gammas);
 | |
| 957 | val ab_prem = foldr1 HOLogic.mk_conj ne_comps; | |
| 958 | ||
| 959 | val REL_pos_distrI_tm = let | |
| 960 | val le_relcomps = map2 mk_leq ab_comps var_Rs; | |
| 961 | val assm = mk_OO [equiv_rel_a, list_comb (rel_F_ab, var_Ps), | |
| 962 | equiv_rel_b, list_comb (rel_F_bc, var_Qs)] equiv_rel_c; | |
| 963 | val concl = mk_OO [equiv_rel_a, list_comb (rel_F_ac, var_Rs)] equiv_rel_c; | |
| 964 | in | |
| 965 | mk_Trueprop_implies | |
| 966 | ([assm $ var_x $ var_y'] @ ne_comps @ le_relcomps, concl $ var_x $ var_y') | |
| 967 | end; | |
| 968 | ||
| 969 | val ab_concl = mk_leq | |
| 970 | (mk_OO [list_comb (rel_F_ab, var_Ps), equiv_rel_b] (list_comb (rel_F_bc, var_Qs))) | |
| 971 | (mk_OO [equiv_rel_a, list_comb (rel_F_ac, ab_comps)] (equiv_rel_c)); | |
| 972 | val ab_imp = Logic.mk_implies (apply2 HOLogic.mk_Trueprop (ab_prem, ab_concl)); | |
| 973 | val rel_pos_distr = fold_rev Logic.all (var_Ps @ var_Qs) ab_imp; | |
| 974 | ||
| 975 |     (* {(x, y) . REL x y} *)
 | |
| 976 | fun mk_rel_pairs rel = mk_case_prod (var_x, var_x') (rel $ var_x $ var_x') | |
| 977 | val rel_pairs = mk_rel_pairs equiv_rel_a; | |
| 978 | ||
| 979 |     (* rel_Inter: \<And>S. \<lbrakk> S \<noteq> {}; \<Inter>S \<noteq> {} \<rbrakk> \<Longrightarrow>
 | |
| 980 |       (\<Inter>A\<in>S. {(x, y). REL x y} `` {x. set_F x \<subseteq> A}) \<subseteq> {(x, y). REL x y} `` {x. set_F x \<subseteq> \<Inter>S} *)
 | |
| 981 | fun rel_Inter_from_set_F (var_A, var_S) set_F = let | |
| 982 | ||
| 983 | val typ_aset = fastype_of var_A; | |
| 984 | ||
| 985 | (* \<Inter>S *) | |
| 986 | val inter_S = Inf_const typ_aset $ var_S; | |
| 987 | ||
| 988 |       (* [S \<noteq> {}, \<Inter>S \<noteq> {}] *)
 | |
| 989 | fun not_empty x = let val ty = fastype_of x | |
| 990 | in HOLogic.mk_not (HOLogic.mk_eq (x, bot_const ty)) end; | |
| 991 | val prems = map (HOLogic.mk_Trueprop o not_empty) [var_S, inter_S]; | |
| 992 | ||
| 993 |       (* {x. set_F x \<subseteq> A} *)
 | |
| 994 | val setF_sub_A = mk_in [var_A] [set_F] typ_aF; | |
| 995 | ||
| 996 |       (* {x. set_F x \<subseteq> \<Inter>S} *)
 | |
| 997 | val setF_sub_S = mk_in [inter_S] [set_F] typ_aF; | |
| 998 | ||
| 999 | val lhs = Inf_const (HOLogic.mk_setT typ_aF) $ (mk_image | |
| 1000 | (absfree (dest_Free var_A) (Image_const typ_aF $ rel_pairs $ setF_sub_A)) $ var_S); | |
| 1001 | val rhs = Image_const typ_aF $ rel_pairs $ setF_sub_S; | |
| 1002 | val concl = HOLogic.mk_Trueprop (mk_leq lhs rhs); | |
| 1003 | ||
| 1004 | in Logic.all var_S (Logic.list_implies (prems, concl)) end; | |
| 1005 | ||
| 1006 | val rel_Inters = map2 rel_Inter_from_set_F (var_As ~~ var_Ss) sets_F; | |
| 1007 | ||
| 1008 | (* map_F_Just = map_F Just *) | |
| 1009 | val map_F_Just = let | |
| 1010 | val option_tys = map mk_MaybeT alphas; | |
| 1011 | val somes = map Just_const alphas; | |
| 1012 | in list_comb (subst_atomic_types (betas ~~ option_tys) map_F, somes) end; | |
| 1013 | ||
| 1014 | fun mk_set_F'_tm typ_a set_F = | |
| 1015 | let | |
| 1016 | val typ_aset = HOLogic.mk_setT typ_a; | |
| 1017 | ||
| 1018 |         (* set_F' x = (\<Inter>y\<in>{y. REL (map_F Just x) y}. UNION (set_F y) set_Maybe) *)
 | |
| 1019 | val sbind = mk_UNION (subst_Maybe alphas set_F $ var_mx) (set_Maybe_const typ_a); | |
| 1020 | val collection = HOLogic.Collect_const typ_MaybeF $ absfree (dest_Free var_mx) | |
| 1021 | (subst_Maybe alphas equiv_rel_a $ (map_F_Just $ var_x) $ var_mx); | |
| 1022 | val set_F'_tm = lambda var_x | |
| 1023 | (Inf_const typ_aset $ (mk_image (absfree (dest_Free var_mx) sbind) $ collection)); | |
| 1024 | in | |
| 1025 | set_F'_tm | |
| 1026 | end | |
| 1027 | ||
| 1028 | val sets = mk_sets_of_bnf (replicate live deads) (replicate live alphas) bnf_F; | |
| 1029 | val sets' = map2 mk_set_F'_tm alphas sets; | |
| 1030 | ||
| 1031 | val (Iwits, wit_goals) = | |
| 1032 | prepare_wits true repT wits opts alphas wits_F var_as var_as' sets' lthy; | |
| 1033 | ||
| 71354 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1034 | val goals = rel_pos_distr :: rel_Inters @ | 
| 71262 | 1035 | (case wits of NONE => [] | _ => wit_goals); | 
| 1036 | ||
| 1037 | val plugins = | |
| 1038 | get_first (fn Plugins_Option f => SOME (f lthy) | _ => NONE) (rev opts) |> | |
| 1039 | the_default Plugin_Name.default_filter; | |
| 1040 | ||
| 1041 | fun after_qed thmss lthy = | |
| 71354 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1042 | (case thmss of [rel_pos_distr_thm0] :: thmss => | 
| 71262 | 1043 | let | 
| 1044 | val equiv_thm' = REL_rewr_all lthy equiv_thm; | |
| 1045 | val rel_pos_distr_thm = | |
| 1046 |             @{thm equivp_add_relconj} OF [equiv_thm', equiv_thm', rel_pos_distr_thm0];
 | |
| 1047 | ||
| 1048 | val (rel_Inter_thms, wit_thmss) = apply2 (fn f => flat (f live thmss)) (take, drop); | |
| 1049 | ||
| 1050 | (* construct map_G, sets_G, bd_G, rel_G and wits_G *) | |
| 1051 | ||
| 1052 | (* map_G f = abs_G o map_F f o rep_G *) | |
| 1053 | val map_G = fold_rev lambda var_fs (HOLogic.mk_comp (HOLogic.mk_comp | |
| 1054 | (subst_Ts abs_G betas, list_comb (map_F, var_fs)), rep_G)); | |
| 1055 | val map_raw = fold_rev lambda var_fs (list_comb (map_F, var_fs)) | |
| 1056 | |> subst_atomic_types (betas ~~ gammas); | |
| 1057 | ||
| 1058 | (* Define set_G and the three auxiliary definitions (set_F', F_in, F_in') *) | |
| 1059 | fun mk_set_G var_A set_F lthy = let | |
| 1060 | val typ_a = HOLogic.dest_setT (fastype_of var_A); | |
| 1061 | val set_F'_tm = mk_set_F'_tm typ_a set_F | |
| 1062 | ||
| 1063 | val (set_F', lthy) = define lthy (suffix set_F "'") 1 set_F'_tm; | |
| 1064 | ||
| 1065 | (* set_G = set_F' o rep_G *) | |
| 1066 | val set_G = HOLogic.mk_comp (#tm set_F', rep_G); | |
| 1067 | ||
| 1068 |               (* F_in A = {x. set_F x \<subseteq> A} *)
 | |
| 1069 | val F_in_tm = lambda var_A (mk_in [var_A] [set_F] typ_aF); | |
| 1070 | val (F_in, lthy) = define lthy (suffix set_F "_in") 1 F_in_tm; | |
| 1071 | ||
| 1072 |               (* F_in' A = map_F Inr -` ({(x, y). REL x y} `` F_in (insert (Inl ()) (Inr ` A))) *)
 | |
| 1073 | val F_in' = lambda var_A (mk_vimage map_F_Just (Image_const typ_MaybeF $ | |
| 1074 | subst_Maybe alphas rel_pairs $ (subst_Maybe alphas (#tm F_in) $ mk_insert | |
| 1075 | (mk_Nothing typ_a) (mk_image (Just_const typ_a) $ var_A)))); | |
| 1076 | val (F_in', lthy) = define lthy (suffix set_F "_in'") 1 F_in'; | |
| 1077 | ||
| 1078 |             in ((set_G, {set_F'=set_F', F_in=F_in, F_in'=F_in'}), lthy) end;
 | |
| 1079 | ||
| 1080 | val ((sets_G, set_F'_aux_defs), lthy) = | |
| 1081 |             @{fold_map 2} mk_set_G var_As sets_F lthy |>> split_list;
 | |
| 1082 | ||
| 1083 | (* bd_G = bd_F *) | |
| 1084 | val bd_G = mk_bd_of_bnf deads alphas bnf_F; | |
| 1085 | ||
| 1086 | (* rel_F' A = | |
| 1087 | BNF_Def.vimage2p (map_F Just) (map_F Just) ((\<cong>) OO rel_F (rel_Maybe A) OO (\<cong>)) *) | |
| 1088 |           val rel_Maybes = @{map 3} (fn v => fn aT => fn bT => rel_Maybe_const aT bT $ v);
 | |
| 1089 | val rel_F'_tm = let val equiv_equiv_rel_option = subst_Ts equiv_rel_a' o map mk_MaybeT in | |
| 1090 | mk_vimage2p map_F_Just (subst_atomic_types (alphas ~~ betas) map_F_Just) $ | |
| 1091 | mk_OO [equiv_equiv_rel_option alphas, list_comb (rel_F_option, rel_Maybes var_Ps alphas betas)] | |
| 1092 | (equiv_equiv_rel_option betas) end; | |
| 1093 | ||
| 1094 | val (rel_F', lthy) = | |
| 1095 | define lthy (suffix rel_F_ab "'") (live+2) (fold_rev lambda var_Ps rel_F'_tm); | |
| 1096 | ||
| 1097 | (* rel_G A = vimage2p rep_G rep_G (rel_F' A) *) | |
| 1098 | val rel_G = fold_rev lambda var_Ps (mk_vimage2p rep_G (subst_Ts rep_G betas) $ rel_F'_tm); | |
| 1099 | val rel_raw = fold_rev lambda var_Ps rel_F'_tm | |
| 1100 | |> subst_atomic_types (betas ~~ gammas); | |
| 1101 | ||
| 1102 | (* auxiliary lemmas *) | |
| 1103 | val bd_card_order = bd_card_order_of_bnf bnf_F; | |
| 1104 | val bd_cinfinite = bd_cinfinite_of_bnf bnf_F; | |
| 75624 | 1105 | val bd_regularCard = bd_regularCard_of_bnf bnf_F; | 
| 71262 | 1106 | val in_rel = in_rel_of_bnf bnf_F; | 
| 1107 | val map_F_comp = map_comp_of_bnf bnf_F; | |
| 1108 | val map_F_comp0 = map_comp0_of_bnf bnf_F; | |
| 1109 | val map_F_cong = map_cong_of_bnf bnf_F; | |
| 1110 | val map_F_id0 = map_id0_of_bnf bnf_F; | |
| 1111 | val map_F_id = map_id_of_bnf bnf_F; | |
| 1112 | val rel_conversep = rel_conversep_of_bnf bnf_F; | |
| 71354 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1113 | val rel_flip = rel_flip_of_bnf bnf_F; | 
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1114 | val rel_eq_onp = rel_eq_onp_of_bnf bnf_F; | 
| 71262 | 1115 | val rel_Grp = rel_Grp_of_bnf bnf_F; | 
| 1116 | val rel_OO = rel_OO_of_bnf bnf_F; | |
| 1117 | val rel_map = rel_map_of_bnf bnf_F; | |
| 1118 | val rel_refl_strong = rel_refl_strong_of_bnf bnf_F; | |
| 1119 | val set_bd_thms = set_bd_of_bnf bnf_F; | |
| 1120 | val set_map_thms = set_map_of_bnf bnf_F; | |
| 1121 | ||
| 71354 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1122 | |
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1123 | |
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1124 |           (* map_F_respect: @{term "((=) ===> REL ===> REL) map_F map_F"} *)
 | 
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1125 | val map_F_respect = HOLogic.mk_Trueprop (fold_rev mk_rel_fun (map2 (fn xT => fn yT => | 
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1126 | HOLogic.eq_const (xT --> yT)) alphas betas @ [equiv_rel_a]) (equiv_rel_b) $ map_F $ map_F); | 
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1127 | |
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1128 | fun map_F_respect_tac ctxt = | 
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1129 | HEADGOAL (EVERY' | 
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1130 |              [REPEAT_DETERM_N (live + 1) o rtac ctxt @{thm rel_funI}, hyp_subst_tac ctxt,
 | 
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1131 | rtac ctxt (BNF_FP_Util.split_conj_prems live rel_pos_distr_thm0 OF | 
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1132 |                 replicate live @{thm Grp_conversep_nonempty} RS rev_mp),
 | 
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1133 |               rtac ctxt impI, dtac ctxt @{thm predicate2D}, etac ctxt @{thm relcomppI2[rotated]},
 | 
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1134 |               rtac ctxt (rel_Grp RS @{thm predicate2_eqD} RS iffD2), rtac ctxt @{thm GrpI[OF refl]},
 | 
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1135 | REPEAT_DETERM o resolve_tac ctxt [CollectI, conjI, subset_UNIV], | 
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1136 | rtac ctxt (rel_flip RS iffD2), | 
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1137 |               rtac ctxt (rel_Grp RS @{thm predicate2_eqD} RS iffD2), rtac ctxt @{thm GrpI[OF refl]},
 | 
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1138 | REPEAT_DETERM o resolve_tac ctxt [CollectI, conjI, subset_UNIV], | 
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1139 |               SELECT_GOAL (unfold_thms_tac ctxt (rel_eq_onp :: @{thms Grp_conversep_eq_onp})),
 | 
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1140 |               etac ctxt @{thm predicate2D[OF rel_conj_eq_onp, rotated]},
 | 
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1141 | rtac ctxt equiv_thm']); | 
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1142 | |
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1143 | val map_F_respect_thm = prove lthy [] map_F_respect map_F_respect_tac; | 
| 
c71a44893645
eliminated one redundant proof obligation in lift_bnf for quotients
 traytel parents: 
71262diff
changeset | 1144 | |
| 71262 | 1145 | val rel_funD = mk_rel_funDN (live+1); | 
| 1146 | val map_F_rsp = (rel_funD map_F_respect_thm) OF (replicate live refl); | |
| 1147 | fun map_F_rsp_of tms ctxt = (infer_instantiate' ctxt (NONE :: NONE | |
| 1148 | :: map (SOME o Thm.cterm_of ctxt) tms) map_F_rsp) | |
| 1149 | ||
| 1150 | val qthms = let | |
| 1151 | fun equivp_THEN thm = REL_rewr_all lthy equiv_thm RS thm; | |
| 1152 | fun Quotient3_THEN thm = REL_rewr_all lthy quot_thm RS thm; | |
| 1153 | in | |
| 1154 |               {abs_rep = Quotient3_THEN @{thm Quotient3_abs_rep},
 | |
| 1155 |                rel_abs = Quotient3_THEN @{thm Quotient3_rel_abs},
 | |
| 1156 |                rep_abs = Quotient3_THEN @{thm Quotient3_rep_abs},
 | |
| 1157 |                rep_reflp = Quotient3_THEN @{thm Quotient3_rep_reflp},
 | |
| 1158 |                rep_abs_rsp = Quotient3_THEN @{thm rep_abs_rsp},
 | |
| 1159 |                reflp = equivp_THEN @{thm equivp_reflp},
 | |
| 1160 |                symp = equivp_THEN @{thm equivp_symp},
 | |
| 1161 |                transp = equivp_THEN @{thm equivp_transp},
 | |
| 1162 | REL = REL_def} | |
| 1163 | end; | |
| 1164 | ||
| 1165 | (* lemma REL_OO_REL_left: REL OO REL OO R = REL OO R *) | |
| 1166 | val REL_OO_REL_left_thm = let | |
| 1167 | val tm = mk_Trueprop_eq | |
| 1168 | (mk_OO [equiv_rel_a, equiv_rel_a] var_R, mk_relcompp equiv_rel_a var_R) | |
| 1169 | fun tac ctxt = HEADGOAL (EVERY' | |
| 1170 | [rtac ctxt ext, | |
| 1171 | rtac ctxt ext, | |
| 1172 | rtac ctxt iffI, | |
| 1173 |                 TWICE (etac ctxt @{thm relcomppE}),
 | |
| 1174 |                 rtac ctxt @{thm relcomppI},
 | |
| 1175 | etac ctxt (#transp qthms), | |
| 1176 | TWICE (assume_tac ctxt), | |
| 1177 |                 etac ctxt @{thm relcomppE},
 | |
| 1178 |                 etac ctxt @{thm relcomppI},
 | |
| 1179 |                 rtac ctxt @{thm relcomppI},
 | |
| 1180 | rtac ctxt (#reflp qthms), | |
| 1181 | assume_tac ctxt]); | |
| 1182 | in prove lthy [var_R] tm tac end; | |
| 1183 | ||
| 1184 | (* Generate theorems related to the setters *) | |
| 1185 | val map_F_fs = list_comb (map_F, var_fs); | |
| 1186 | ||
| 1187 | (* aset aset asetset bset typ_b typ_b *) | |
| 1188 | fun mk_set_F'_thmss (((((var_A, var_A'), var_S), var_B), var_b), var_b') | |
| 1189 |                 set_F {set_F', F_in, F_in'} rel_Inter_thm set_map_F_thm (idx, vf) =
 | |
| 1190 | let | |
| 1191 | val (var_f, var_fs') = case vf of | |
| 1192 | (f :: fs) => (f, fs) | |
| 1193 | | _ => error "won't happen"; | |
| 1194 | ||
| 1195 | val typ_a = fastype_of var_f |> dest_funT |> fst; | |
| 1196 | val typ_b = fastype_of var_b; | |
| 1197 | val (typ_asetset, typ_aset) = `HOLogic.mk_setT (fastype_of var_A); | |
| 1198 | ||
| 1199 | val map_F_fs_x = map_F_fs $ var_x; | |
| 1200 | ||
| 1201 | (* F_in'_mono: A \<subseteq> B \<Longrightarrow> F_in' A \<subseteq> F_in' B *) | |
| 1202 | val F_in'_mono_tm = mk_Trueprop_implies | |
| 1203 | ([mk_leq var_A var_A'], mk_leq (#tm F_in' $ var_A) (#tm F_in' $ var_A')); | |
| 1204 | fun F_in'_mono_tac ctxt = | |
| 1205 | unfold_thms_tac ctxt [#def F_in', #def F_in] THEN | |
| 1206 | HEADGOAL (EVERY' | |
| 1207 | [rtac ctxt subsetI, | |
| 1208 | etac ctxt vimageE, | |
| 1209 |                   etac ctxt @{thm ImageE},
 | |
| 1210 | etac ctxt CollectE, | |
| 1211 | etac ctxt CollectE, | |
| 1212 |                   dtac ctxt @{thm case_prodD},
 | |
| 1213 | hyp_subst_tac ctxt, | |
| 1214 | rtac ctxt (vimageI OF [refl]), | |
| 1215 |                   rtac ctxt @{thm ImageI},
 | |
| 1216 | rtac ctxt CollectI, | |
| 1217 |                   rtac ctxt @{thm case_prodI},
 | |
| 1218 | assume_tac ctxt ORELSE' rtac ctxt refl, | |
| 1219 | rtac ctxt CollectI, | |
| 1220 | etac ctxt subset_trans, | |
| 1221 |                   etac ctxt (@{thm insert_mono} OF [@{thm image_mono}])]);
 | |
| 1222 | val F_in'_mono_thm = prove lthy [var_A, var_A'] F_in'_mono_tm F_in'_mono_tac; | |
| 1223 | ||
| 1224 | (* F_in'_Inter: F_in' (\<Inter>S) = (\<Inter>A\<in>S. F_in' A) *) | |
| 1225 | val F_in'_Inter_tm = mk_Trueprop_eq ((#tm F_in' $ (Inf_const typ_aset $ var_S)), | |
| 1226 | (Inf_const (HOLogic.mk_setT typ_aF) $ (mk_image (#tm F_in') $ var_S))); | |
| 1227 | fun F_in'_Inter_tac ctxt = | |
| 1228 | Local_Defs.unfold_tac ctxt [#def F_in', #def F_in] | |
| 1229 | THEN HEADGOAL (rtac ctxt (infer_instantiate' ctxt | |
| 1230 |                   [SOME (Thm.cterm_of ctxt (HOLogic.mk_eq (var_S, bot_const typ_asetset)))] @{thm case_split})
 | |
| 1231 | THEN' EVERY' [ | |
| 1232 | hyp_subst_tac ctxt, | |
| 1233 | SELECT_GOAL | |
| 1234 |                       (unfold_thms_tac ctxt @{thms Inter_empty INT_empty UNIV_sum_unit_conv}),
 | |
| 1235 |                     rtac ctxt @{thm set_eqI},
 | |
| 1236 | rtac ctxt iffI, | |
| 1237 | rtac ctxt UNIV_I, | |
| 1238 | rtac ctxt (vimageI OF [refl]), | |
| 1239 |                     rtac ctxt @{thm ImageI},
 | |
| 1240 | rtac ctxt CollectI, | |
| 1241 |                     rtac ctxt @{thm case_prodI},
 | |
| 1242 | rtac ctxt (#reflp qthms), | |
| 1243 | rtac ctxt CollectI, | |
| 1244 | rtac ctxt subset_UNIV, | |
| 1245 |                     etac ctxt @{thm exE[OF ex_in_conv[THEN iffD2]]},
 | |
| 1246 |                     EqSubst.eqsubst_tac ctxt [0] @{thms image_INT[of _ UNIV _ id, simplified]},
 | |
| 1247 |                     rtac ctxt @{thm inj_Inr},
 | |
| 1248 | assume_tac ctxt, | |
| 1249 |                     SELECT_GOAL (unfold_thms_tac ctxt @{thms INT_extend_simps vimage_INT[symmetric]}),
 | |
| 1250 |                     rtac ctxt @{thm arg_cong2[where f=vimage, OF refl]},
 | |
| 1251 | rtac ctxt equalityI, | |
| 1252 | rtac ctxt subsetI, | |
| 1253 |                     rtac ctxt @{thm InterI},
 | |
| 1254 | etac ctxt imageE, | |
| 1255 |                     etac ctxt @{thm ImageE},
 | |
| 1256 | etac ctxt CollectE, | |
| 1257 | etac ctxt CollectE, | |
| 1258 |                     dtac ctxt @{thm case_prodD},
 | |
| 1259 | hyp_subst_tac ctxt, | |
| 1260 |                     rtac ctxt @{thm ImageI[OF CollectI]},
 | |
| 1261 |                     etac ctxt @{thm case_prodI} ORELSE' (SELECT_GOAL
 | |
| 1262 |                       (unfold_thms_tac ctxt @{thms prod.case}) THEN' rtac ctxt @{thm refl}),
 | |
| 1263 | rtac ctxt CollectI, | |
| 1264 | etac ctxt subset_trans, | |
| 1265 |                     etac ctxt @{thm INT_lower[OF imageI]},
 | |
| 1266 |                     rtac ctxt (@{thm subset_trans} OF [asm_rl, rel_Inter_thm]),
 | |
| 1267 |                     K (unfold_thms_tac ctxt @{thms image_image}),
 | |
| 1268 | rtac ctxt subset_refl, | |
| 1269 |                     K (unfold_thms_tac ctxt @{thms INT_extend_simps ex_in_conv[symmetric]}),
 | |
| 1270 | rtac ctxt exI, | |
| 1271 | rtac ctxt imageI, | |
| 1272 | assume_tac ctxt, | |
| 1273 | rtac ctxt exI, | |
| 1274 |                     rtac ctxt @{thm InterI},
 | |
| 1275 | etac ctxt imageE, | |
| 1276 | hyp_subst_tac ctxt, | |
| 1277 |                     rtac ctxt @{thm insertI1}]);
 | |
| 1278 | ||
| 1279 | val F_in'_Inter_thm = prove lthy [var_S] F_in'_Inter_tm F_in'_Inter_tac; | |
| 1280 | ||
| 1281 | (* set_F'_respect: (REL ===> (=)) set_F' set_F' *) | |
| 1282 | val set_F'_respect_tm = HOLogic.mk_Trueprop (mk_rel_fun equiv_rel_a | |
| 1283 | (HOLogic.eq_const typ_aset) $ #tm set_F' $ #tm set_F'); | |
| 1284 |               fun set_F'_respect_tac ctxt = unfold_thms_tac ctxt (#def set_F' :: @{thms rel_fun_def})
 | |
| 1285 | THEN HEADGOAL (EVERY' | |
| 1286 | [TWICE (rtac ctxt allI), | |
| 1287 | rtac ctxt impI, | |
| 1288 | dtac ctxt (map_F_rsp_of (map Just_const alphas) ctxt), | |
| 1289 |                   rtac ctxt @{thm INF_cong},
 | |
| 1290 |                   rtac ctxt @{thm Collect_eqI},
 | |
| 1291 | rtac ctxt iffI, | |
| 1292 | etac ctxt (#transp qthms OF [#symp qthms]), | |
| 1293 | assume_tac ctxt, | |
| 1294 | etac ctxt (#transp qthms), | |
| 1295 | assume_tac ctxt, | |
| 1296 | rtac ctxt refl]); | |
| 1297 | ||
| 1298 |               (* F_in'_alt2: F_in' A = {x. set_F' x \<subseteq> A} *)
 | |
| 1299 | val F_in'_alt2_tm = mk_Trueprop_eq | |
| 1300 | (#tm F_in' $ var_A, mk_in [var_A] [#tm set_F'] typ_aF); | |
| 1301 | fun F_in'_alt2_tac ctxt = HEADGOAL (rtac ctxt equalityI THEN' | |
| 1302 | (Subgoal.FOCUS o K) (unfold_thms_tac ctxt (map #def [set_F', F_in', F_in])) ctxt | |
| 1303 | THEN' EVERY' | |
| 1304 | [rtac ctxt subsetI, | |
| 1305 | rtac ctxt CollectI, | |
| 1306 | rtac ctxt subsetI, | |
| 1307 | dtac ctxt vimageD, | |
| 1308 |                   etac ctxt @{thm ImageE},
 | |
| 1309 | etac ctxt CollectE, | |
| 1310 | etac ctxt CollectE, | |
| 1311 |                   dtac ctxt @{thm case_prodD},
 | |
| 1312 |                   dtac ctxt @{thm InterD},
 | |
| 1313 |                   rtac ctxt @{thm imageI[OF CollectI]},
 | |
| 1314 | etac ctxt (#symp qthms), | |
| 1315 |                   etac ctxt @{thm UnionE},
 | |
| 1316 | etac ctxt imageE, | |
| 1317 | hyp_subst_tac ctxt, | |
| 1318 |                   etac ctxt @{thm subset_lift_sum_unitD},
 | |
| 1319 |                   etac ctxt @{thm setr.cases},
 | |
| 1320 | hyp_subst_tac ctxt, | |
| 1321 | assume_tac ctxt]) | |
| 1322 | THEN unfold_thms_tac ctxt [#def set_F'] THEN | |
| 1323 | (HEADGOAL o EVERY') | |
| 1324 | [rtac ctxt subsetI, | |
| 1325 | etac ctxt CollectE, | |
| 1326 | etac ctxt (subsetD OF [F_in'_mono_thm]), | |
| 1327 | EqSubst.eqsubst_tac ctxt [0] [F_in'_Inter_thm], | |
| 1328 |                   rtac ctxt @{thm InterI}] THEN
 | |
| 1329 |                 REPEAT_DETERM (HEADGOAL (etac ctxt @{thm imageE} THEN' hyp_subst_tac ctxt)) THEN
 | |
| 1330 | (HEADGOAL o EVERY') | |
| 1331 | [etac ctxt CollectE, | |
| 1332 | SELECT_GOAL (unfold_thms_tac ctxt (map #def [F_in', F_in])), | |
| 1333 |                   rtac ctxt @{thm vimageI[OF refl]},
 | |
| 1334 |                   rtac ctxt @{thm ImageI},
 | |
| 1335 | rtac ctxt CollectI, | |
| 1336 |                   rtac ctxt @{thm case_prodI},
 | |
| 1337 | etac ctxt (#symp qthms), | |
| 1338 | rtac ctxt CollectI, | |
| 1339 | rtac ctxt subsetI, | |
| 1340 |                   rtac ctxt @{thm sum_insert_Inl_unit},
 | |
| 1341 | assume_tac ctxt, | |
| 1342 | hyp_subst_tac ctxt, | |
| 1343 | rtac ctxt imageI, | |
| 1344 |                   rtac ctxt @{thm UnionI},
 | |
| 1345 | rtac ctxt imageI, | |
| 1346 | assume_tac ctxt, | |
| 1347 |                   rtac ctxt @{thm setr.intros[OF refl]}];
 | |
| 1348 | val F_in'_alt2_thm = prove lthy [var_A] F_in'_alt2_tm F_in'_alt2_tac; | |
| 1349 | ||
| 1350 |               (* set_F'_alt: set_F' x = \<Inter>{A. x \<in> F_in' A} *)
 | |
| 1351 | val set_F'_alt_tm = mk_Trueprop_eq (#tm set_F' $ var_x, | |
| 1352 | Inf_const typ_aset $ mk_Collect (var_A, HOLogic.mk_mem (var_x, #tm F_in' $ var_A))); | |
| 1353 | fun set_F'_alt_tac ctxt = unfold_thms_tac ctxt [F_in'_alt2_thm] | |
| 1354 | THEN HEADGOAL (EVERY' | |
| 1355 |                   [rtac ctxt @{thm set_eqI},
 | |
| 1356 | rtac ctxt iffI, | |
| 1357 |                   rtac ctxt @{thm InterI},
 | |
| 1358 | etac ctxt CollectE, | |
| 1359 | etac ctxt CollectE, | |
| 1360 | dtac ctxt subsetD, | |
| 1361 | assume_tac ctxt, | |
| 1362 | assume_tac ctxt, | |
| 1363 |                   etac ctxt @{thm InterD},
 | |
| 1364 | rtac ctxt CollectI, | |
| 1365 | rtac ctxt CollectI, | |
| 1366 | rtac ctxt subset_refl]); | |
| 1367 | val set_F'_alt_thm = prove lthy [var_x] set_F'_alt_tm set_F'_alt_tac; | |
| 1368 | ||
| 1369 | (* map_F_in_F_in'I: x \<in> F_in' B \<Longrightarrow> map_F f x \<in> F_in' (f ` B) *) | |
| 1370 | val map_F_in_F_in'I_tm = mk_Trueprop_implies ([HOLogic.mk_mem (var_x, #tm F_in' $ var_A')], | |
| 1371 | HOLogic.mk_mem (map_F_fs_x, subst_b (#tm F_in') $ (mk_image var_f $ var_A'))); | |
| 1372 | fun map_F_in_F_in'I_tac ctxt = | |
| 1373 |                 Local_Defs.unfold_tac ctxt ([#def F_in', #def F_in, Bex_def] @ @{thms vimage_def Image_iff}) THEN
 | |
| 1374 | HEADGOAL (EVERY' | |
| 1375 |                   [etac ctxt @{thm CollectE},
 | |
| 1376 | etac ctxt exE, | |
| 1377 | etac ctxt conjE, | |
| 1378 |                   etac ctxt @{thm CollectE},
 | |
| 1379 |                   etac ctxt @{thm CollectE},
 | |
| 1380 |                   dtac ctxt @{thm case_prodD},
 | |
| 1381 |                   rtac ctxt @{thm CollectI},
 | |
| 1382 | rtac ctxt exI, | |
| 1383 |                   rtac ctxt @{thm conjI[rotated]},
 | |
| 1384 |                   rtac ctxt @{thm CollectI},
 | |
| 1385 |                   rtac ctxt @{thm case_prodI},
 | |
| 1386 | dtac ctxt (map_F_rsp_of (map mk_Maybe_map var_fs) ctxt), | |
| 1387 |                   SELECT_GOAL (unfold_thms_tac ctxt (map_F_comp :: @{thms o_def map_sum.simps})),
 | |
| 1388 | assume_tac ctxt, | |
| 1389 | rtac ctxt CollectI, | |
| 1390 | SELECT_GOAL (unfold_thms_tac ctxt set_map_thms), | |
| 1391 |                   etac ctxt @{thm image_map_sum_unit_subset}]);
 | |
| 1392 | val map_F_in_F_in'I_thm = | |
| 1393 | prove lthy (var_A' :: var_x :: var_fs) map_F_in_F_in'I_tm map_F_in_F_in'I_tac; | |
| 1394 | ||
| 1395 |               (* REL_preimage_eq: C \<inter> range f \<noteq> {} \<Longrightarrow>
 | |
| 1396 |                     {(a, b). REL a b} `` {x. set_F x \<subseteq> f -` C} =
 | |
| 1397 |                         map_F f -` {(a, b). REL a b} `` {x. set_F x \<subseteq> C} *)
 | |
| 1398 | val REL_preimage_eq_tm = mk_Trueprop_implies ([HOLogic.mk_not (HOLogic.mk_eq | |
| 1399 |                   (HOLogic.mk_binop @{const_name inf} (var_B, mk_image var_f $ HOLogic.mk_UNIV typ_a),
 | |
| 1400 | bot_const (HOLogic.mk_setT typ_b)))], HOLogic.mk_eq (Image_const typ_aF $ | |
| 1401 | rel_pairs $ mk_in [mk_vimage var_f var_B] [set_F] typ_aF, mk_vimage map_F_fs | |
| 1402 | (Image_const typ_bF $ subst_b rel_pairs $ mk_in [var_B] [subst_b set_F] typ_bF))); | |
| 1403 | ||
| 1404 |               (* Bs \<inter> range fs \<noteq> {} \<Longrightarrow> set_F xb \<subseteq> Bs \<Longrightarrow> REL xb (map_F fs x)
 | |
| 1405 |                     \<Longrightarrow> x \<in> {(x, x'). REL x x'} `` {x. set_F x \<subseteq> fs -` Bs}              *)
 | |
| 1406 |               fun subgoal_tac {context = ctxt, params, ...} = let
 | |
| 1407 | val (x, y) = case params of | |
| 1408 | [(_, x), _, (_, y)] => (x, y) | |
| 1409 | | _ => error "won't happen"; | |
| 1410 | val cond = HOLogic.mk_conj (apply2 HOLogic.mk_mem ((var_b, var_B), (var_b', var_B))); | |
| 1411 | ||
| 1412 | (* ["\<lambda>x y. x \<in> B \<and> y \<in> B", "(Grp UNIV f_1)\<inverse>\<inverse>"] *) | |
| 1413 | val cvars = var_fs |> maps (fn f => let val fT = fastype_of f in | |
| 1414 | map (SOME o Thm.cterm_of ctxt) | |
| 1415 | [if f = var_f then | |
| 1416 | fold_rev lambda [var_b, var_b'] cond else HOLogic.eq_const (range_type fT), | |
| 1417 | mk_conversep (mk_Grp (HOLogic.mk_UNIV (domain_type fT)) f)] end); | |
| 1418 | val rel_pos_distr_thm_inst = infer_instantiate' ctxt (cvars @ [SOME y,SOME x]) | |
| 1419 |                     (@{thm predicate2D} OF [rel_pos_distr_thm]);
 | |
| 1420 | ||
| 1421 | (* GrpI[of "map_F f1 .. fN" x, OF refl CollectI, OF "B1 \<subseteq> UNIV \<and> ... \<and> Bn \<subseteq> UNIV"] *) | |
| 1422 | fun subset_UNIVs n = fold (fn a => fn b => conjI OF [a, b]) (replicate (n-1) | |
| 1423 |                     @{thm subset_UNIV}) @{thm subset_UNIV};
 | |
| 1424 | val GrpI_inst = infer_instantiate' ctxt (map SOME [Thm.cterm_of ctxt map_F_fs, x]) | |
| 1425 |                     @{thm GrpI} OF [refl, CollectI] OF [subset_UNIVs live];
 | |
| 1426 | ||
| 1427 | in EVERY [ | |
| 1428 | HEADGOAL (Method.insert_tac ctxt [rel_pos_distr_thm_inst]), | |
| 1429 | unfold_thms_tac ctxt [rel_conversep, rel_OO, rel_Grp], | |
| 1430 |                   HEADGOAL (etac ctxt @{thm meta_impE}),
 | |
| 1431 |                   REPEAT_DETERM_N (live-1) (HEADGOAL (rtac ctxt @{thm conjI[rotated]})),
 | |
| 1432 |                   REPEAT_DETERM_N live (HEADGOAL (etac ctxt @{thm relcompp_mem_Grp_neq_bot} ORELSE'
 | |
| 1433 |                       rtac ctxt @{thm relcompp_eq_Grp_neq_bot})),
 | |
| 1434 |                   HEADGOAL (EVERY' [etac ctxt @{thm meta_impE},
 | |
| 1435 |                     rtac ctxt @{thm relcomppI},
 | |
| 1436 | rtac ctxt (#reflp qthms), | |
| 1437 |                     rtac ctxt @{thm relcomppI},
 | |
| 1438 | rtac ctxt rel_refl_strong]), | |
| 1439 | REPEAT_DETERM_N idx (HEADGOAL (rtac ctxt refl)), | |
| 1440 | HEADGOAL (rtac ctxt conjI THEN' TWICE (etac ctxt subsetD THEN' assume_tac ctxt)), | |
| 1441 | REPEAT_DETERM_N (live-idx-1) (HEADGOAL (rtac ctxt refl)), | |
| 1442 | HEADGOAL (EVERY' | |
| 1443 |                     [rtac ctxt @{thm relcomppI},
 | |
| 1444 | assume_tac ctxt, | |
| 1445 |                     rtac ctxt @{thm relcomppI},
 | |
| 1446 |                     rtac ctxt @{thm conversepI},
 | |
| 1447 | rtac ctxt GrpI_inst, | |
| 1448 | rtac ctxt (#reflp qthms), | |
| 1449 |                     etac ctxt @{thm relcomppE},
 | |
| 1450 |                     etac ctxt @{thm relcomppE},
 | |
| 1451 |                     etac ctxt @{thm relcomppE},
 | |
| 1452 |                     etac ctxt @{thm conversepE},
 | |
| 1453 |                     etac ctxt @{thm GrpE},
 | |
| 1454 | hyp_subst_tac ctxt, | |
| 1455 |                     rtac ctxt @{thm ImageI},
 | |
| 1456 | rtac ctxt CollectI, | |
| 1457 |                     rtac ctxt @{thm case_prodI},
 | |
| 1458 | assume_tac ctxt, | |
| 1459 | EqSubst.eqsubst_asm_tac ctxt [1] rel_map, | |
| 1460 | EqSubst.eqsubst_asm_tac ctxt [1] [in_rel_of_bnf bnf_F], | |
| 1461 | etac ctxt exE, | |
| 1462 | etac ctxt CollectE, | |
| 1463 | etac ctxt conjE, | |
| 1464 | etac ctxt conjE, | |
| 1465 | etac ctxt CollectE, | |
| 1466 | hyp_subst_tac ctxt, | |
| 1467 | rtac ctxt CollectI]), | |
| 1468 | unfold_thms_tac ctxt set_map_thms, | |
| 1469 | HEADGOAL (rtac ctxt (subsetI OF [vimageI] OF [refl]) THEN' | |
| 1470 |                             etac ctxt @{thm imageE} THEN' hyp_subst_tac ctxt),
 | |
| 1471 | REPEAT_DETERM_N 6 (HEADGOAL (etac ctxt Drule.thin_rl)), | |
| 1472 | REPEAT_DETERM_N (live-1) (HEADGOAL (etac ctxt conjE)), | |
| 1473 | HEADGOAL (EVERY' [dtac ctxt subsetD, assume_tac ctxt, etac ctxt CollectE]), | |
| 1474 |                   unfold_thms_tac ctxt @{thms split_beta},
 | |
| 1475 | HEADGOAL (etac ctxt conjunct2)] end; | |
| 1476 | ||
| 1477 | fun REL_preimage_eq_tac ctxt = HEADGOAL (EVERY' | |
| 1478 |                 [rtac ctxt @{thm set_eqI},
 | |
| 1479 | rtac ctxt iffI, | |
| 1480 |                 etac ctxt @{thm ImageE},
 | |
| 1481 | etac ctxt CollectE, | |
| 1482 | etac ctxt CollectE, | |
| 1483 |                 dtac ctxt @{thm case_prodD},
 | |
| 1484 | rtac ctxt (vimageI OF [refl]), | |
| 1485 |                 rtac ctxt @{thm ImageI},
 | |
| 1486 | rtac ctxt CollectI, | |
| 1487 |                 rtac ctxt @{thm case_prodI},
 | |
| 1488 | etac ctxt map_F_rsp, | |
| 1489 | rtac ctxt CollectI, | |
| 1490 | EqSubst.eqsubst_tac ctxt [0] [set_map_F_thm], | |
| 1491 |                 etac ctxt @{thm subset_vimage_image_subset},
 | |
| 1492 | etac ctxt vimageE, | |
| 1493 |                 etac ctxt @{thm ImageE},
 | |
| 1494 | TWICE (etac ctxt CollectE), | |
| 1495 |                 dtac ctxt @{thm case_prodD},
 | |
| 1496 | hyp_subst_tac ctxt, | |
| 1497 | Subgoal.FOCUS_PARAMS subgoal_tac ctxt]); | |
| 1498 | ||
| 1499 | val REL_preimage_eq_thm = prove lthy (var_B :: var_fs) REL_preimage_eq_tm REL_preimage_eq_tac; | |
| 1500 | ||
| 1501 | (* set_map_F': set_F' (map_F f x) = f ` set_F' x *) | |
| 1502 | val set_map_F'_tm = mk_Trueprop_eq (subst_b (#tm set_F') | |
| 1503 | $ map_F_fs_x, mk_image var_f $ (#tm set_F' $ var_x)); | |
| 1504 | fun set_map_F'_tac ctxt = HEADGOAL (EVERY' | |
| 1505 |                   [rtac ctxt @{thm set_eqI},
 | |
| 1506 | rtac ctxt iffI, | |
| 1507 | EqSubst.eqsubst_asm_tac ctxt [0] [set_F'_alt_thm], | |
| 1508 |                   etac ctxt @{thm InterD},
 | |
| 1509 | rtac ctxt CollectI, | |
| 1510 | rtac ctxt map_F_in_F_in'I_thm, | |
| 1511 | SELECT_GOAL (unfold_thms_tac ctxt [F_in'_alt2_thm]), | |
| 1512 | rtac ctxt CollectI, | |
| 1513 | rtac ctxt subset_refl, | |
| 1514 | SELECT_GOAL (unfold_thms_tac ctxt [set_F'_alt_thm]), | |
| 1515 |                   rtac ctxt @{thm InterI},
 | |
| 1516 | etac ctxt imageE, | |
| 1517 | etac ctxt CollectE, | |
| 1518 | hyp_subst_tac ctxt, | |
| 1519 |                   etac ctxt @{thm vimageD[OF InterD]},
 | |
| 1520 | rtac ctxt CollectI]) THEN | |
| 1521 | (* map_F f x \<in> F_in' X \<Longrightarrow> x \<in> F_in' (f -` X) *) | |
| 1522 |                   HEADGOAL (Subgoal.FOCUS_PARAMS (fn {context = ctxt, params, ...} =>
 | |
| 1523 | let | |
| 1524 | val X = nth params 1 |> snd |> Thm.term_of; | |
| 1525 | val Inr_img = mk_image (Just_const (HOLogic.dest_setT (fastype_of X))) $ X; | |
| 1526 | fun cvars_of ctxt = map (SOME o Thm.cterm_of ctxt); | |
| 1527 | val cut_thm = infer_instantiate' ctxt (cvars_of ctxt [Inr_img, var_f]) | |
| 1528 |                         @{thm insert_Inl_int_map_sum_unit};
 | |
| 1529 | val preimage_thm = infer_instantiate' ctxt (cvars_of ctxt | |
| 1530 | (filter (fn f => var_f <> f) var_fs |> map mk_Maybe_map)) | |
| 1531 | (cut_thm RS REL_preimage_eq_thm); | |
| 1532 | in EVERY [ | |
| 1533 | unfold_thms_tac ctxt (map #def [F_in', F_in] @ preimage_thm :: map_F_comp :: | |
| 1534 |                           @{thms lift_sum_unit_vimage_commute vimage_comp o_def map_sum.simps}),
 | |
| 1535 |                       unfold_thms_tac ctxt [@{thm o_def[symmetric]}, map_F_comp0],
 | |
| 1536 |                       Local_Defs.fold_tac ctxt @{thms vimage_comp},
 | |
| 1537 | HEADGOAL (etac ctxt (vimageI OF [refl]))] end) ctxt); | |
| 1538 | ||
| 1539 | (* set_F'_subset: set_F' x \<subseteq> set_F x *) | |
| 1540 | val set_F'_subset_tm = HOLogic.mk_Trueprop (mk_leq (#tm set_F' $ var_x) (set_F $ var_x)); | |
| 1541 | fun set_F'_subset_tac ctxt = | |
| 1542 | let val int_e_thm = infer_instantiate' ctxt | |
| 1543 |                   (replicate 3 NONE @ [SOME (Thm.cterm_of ctxt (map_F_Just $ var_x))]) @{thm INT_E};
 | |
| 1544 | in HEADGOAL (EVERY' [SELECT_GOAL (Local_Defs.unfold_tac ctxt [#def set_F']), | |
| 1545 | rtac ctxt subsetI, | |
| 1546 | etac ctxt int_e_thm, | |
| 1547 | SELECT_GOAL (unfold_thms_tac ctxt [set_map_F_thm]), | |
| 1548 |                   etac ctxt @{thm UN_E},
 | |
| 1549 | etac ctxt imageE, | |
| 1550 | hyp_subst_tac ctxt, | |
| 1551 |                   SELECT_GOAL (unfold_thms_tac ctxt @{thms sum_set_simps singleton_iff}),
 | |
| 1552 | hyp_subst_tac ctxt, | |
| 1553 | assume_tac ctxt, | |
| 1554 | etac ctxt notE, | |
| 1555 | rtac ctxt CollectI, | |
| 1556 | rtac ctxt (#reflp qthms)]) | |
| 1557 | end; | |
| 1558 | in | |
| 1559 |               ({F_in'_mono = F_in'_mono_thm,
 | |
| 1560 | F_in'_Inter = F_in'_Inter_thm, | |
| 1561 | set_F'_respect = prove lthy [] set_F'_respect_tm set_F'_respect_tac, | |
| 1562 | F_in'_alt2 = F_in'_alt2_thm, | |
| 1563 | set_F'_alt = set_F'_alt_thm, | |
| 1564 | map_F_in_F_in'I = map_F_in_F_in'I_thm, | |
| 1565 | set_map_F' = prove lthy (var_x :: var_fs) set_map_F'_tm set_map_F'_tac, | |
| 1566 | set_F'_subset = prove lthy [var_x] set_F'_subset_tm set_F'_subset_tac, | |
| 1567 | set_F'_def = #def set_F', | |
| 1568 | F_in_def = #def F_in, | |
| 1569 | F_in'_def = #def F_in'}, (idx + 1, var_fs')) | |
| 1570 | end; | |
| 1571 | ||
| 1572 |           val set_F'_thmss = @{fold_map 5} mk_set_F'_thmss
 | |
| 1573 | (var_As ~~ var_As' ~~ var_Ss ~~ var_Bs ~~ var_bs ~~ var_bs') sets_F set_F'_aux_defs | |
| 1574 | rel_Inter_thms set_map_thms (0, var_fs) | |
| 1575 | |> fst; | |
| 1576 | ||
| 1577 | (* F_in'D: x \<in> F_in' A \<Longrightarrow> \<forall>a\<in>A. f a = g a \<Longrightarrow> REL (map_F f x) (map_F g x) *) | |
| 1578 | fun rel_map_F_fs_map_F_gs subst fs gs = subst equiv_rel_b $ | |
| 1579 | (list_comb (subst map_F, fs) $ var_x) $ (list_comb (subst map_F, gs) $ var_x); | |
| 1580 | val F_in'D_thm = let | |
| 1581 |               fun mk_prem var_a var_Aset {F_in', ...} var_f var_g =
 | |
| 1582 | [HOLogic.mk_mem (var_x, #tm F_in' $ var_Aset), mk_Ball var_Aset | |
| 1583 | ((absfree (dest_Free var_a)) (HOLogic.mk_eq (var_f $ var_a, var_g $ var_a)))]; | |
| 1584 |               val prems = @{map 5} mk_prem var_as var_As set_F'_aux_defs var_fs' var_gs';
 | |
| 1585 | val F_in'D_tm = mk_Trueprop_implies (flat prems, | |
| 1586 | rel_map_F_fs_map_F_gs (subst_Maybe betas) var_fs' var_gs'); | |
| 1587 | ||
| 1588 | fun map_F_rsp_of_case_sums fs ctxt = map_F_rsp_of | |
| 1589 |                 (@{map 2} (fn f => fn T => BNF_FP_Util.mk_case_sum
 | |
| 1590 | (Term.absdummy HOLogic.unitT (mk_Nothing T), f)) fs betas) ctxt; | |
| 1591 | ||
| 1592 | fun mk_var_fgs n = take n var_gs' @ drop n var_fs'; | |
| 1593 | fun F_in'D_tac ctxt = EVERY | |
| 1594 | (unfold_thms_tac ctxt | |
| 1595 |                   (maps (fn {F_in'_def, F_in_def, ...} => [F_in'_def, F_in_def]) set_F'_thmss) ::
 | |
| 1596 | map (REPEAT_DETERM_N live o HEADGOAL) | |
| 1597 | [etac ctxt vimageE, | |
| 1598 |                   etac ctxt @{thm ImageE},
 | |
| 1599 | etac ctxt CollectE THEN' etac ctxt CollectE, | |
| 1600 |                   dtac ctxt @{thm case_prodD}] @
 | |
| 1601 | HEADGOAL (hyp_subst_tac ctxt THEN' rotate_tac (~live)) :: | |
| 1602 | map (fn i => (HEADGOAL o EVERY') | |
| 1603 | [if i < live then rtac ctxt (#transp qthms) else K all_tac, | |
| 1604 | Ctr_Sugar_Tactics.select_prem_tac ctxt live (dresolve_tac ctxt [asm_rl]) i, | |
| 1605 | rtac ctxt (#transp qthms), | |
| 1606 | dtac ctxt (map_F_rsp_of_case_sums (mk_var_fgs (i-1)) ctxt), | |
| 1607 |                   SELECT_GOAL (unfold_thms_tac ctxt (map_F_comp :: @{thms case_sum_o_inj(2)})),
 | |
| 1608 | etac ctxt (#symp qthms), | |
| 1609 | dtac ctxt (map_F_rsp_of_case_sums (mk_var_fgs i) ctxt), | |
| 1610 |                   SELECT_GOAL (unfold_thms_tac ctxt (map_F_comp :: @{thms case_sum_o_inj(2)})),
 | |
| 1611 | EqSubst.eqsubst_tac ctxt [0] [map_F_cong OF replicate i refl @ asm_rl :: replicate (live-i) refl], | |
| 1612 |                   rtac ctxt @{thm sum.case_cong[OF refl refl]},
 | |
| 1613 | etac ctxt bspec, | |
| 1614 | hyp_subst_tac ctxt, | |
| 1615 |                   etac ctxt @{thm subset_lift_sum_unitD},
 | |
| 1616 | assume_tac ctxt, | |
| 1617 | assume_tac ctxt]) (1 upto live)); | |
| 1618 | ||
| 1619 | in prove lthy (var_x :: var_As @ var_fs' @ var_gs') F_in'D_tm F_in'D_tac end; | |
| 1620 | ||
| 1621 | (* map_F_cong': (\<And>a. a \<in> set_F' x \<Longrightarrow> f a = g a) \<Longrightarrow> REL (map_F f x) (map_F g x) *) | |
| 1622 | val map_F_cong'_thm = let | |
| 1623 |               fun mk_prem {set_F', ...} var_a var_f var_g = Logic.all var_a
 | |
| 1624 | (mk_Trueprop_implies ([HOLogic.mk_mem (var_a, #tm set_F' $ var_x)], | |
| 1625 | HOLogic.mk_eq (var_f $ var_a, var_g $ var_a))); | |
| 1626 | val map_F_cong'_tm = Logic.list_implies | |
| 1627 |                 (@{map 4} mk_prem set_F'_aux_defs var_as var_fs var_gs, HOLogic.mk_Trueprop
 | |
| 1628 | (rel_map_F_fs_map_F_gs I var_fs var_gs)); | |
| 1629 | fun Just_o_fun bT f = HOLogic.mk_comp (Just_const bT, f); | |
| 1630 | fun map_F_Just_o_funs fs = list_comb | |
| 1631 | (subst_Maybe betas map_F, map2 Just_o_fun betas fs) $ var_x; | |
| 1632 | fun map_F_cong'_tac ctxt = let | |
| 1633 | val map_F_respect_inst = map_F_rsp | |
| 1634 | |> infer_instantiate' ctxt (map (SOME o Thm.cterm_of ctxt) | |
| 1635 | (map map_F_Just_o_funs [var_fs, var_gs] @ map fromJust_const betas)) | |
| 1636 |                     |> Local_Defs.unfold ctxt (map_F_comp :: @{thms o_assoc
 | |
| 1637 | Fun.o_apply[where f=projr and g=Inr, unfolded sum.sel] id_def[symmetric]}) | |
| 1638 |                     |> Local_Defs.unfold ctxt @{thms id_comp};
 | |
| 1639 | in | |
| 1640 | HEADGOAL (rtac ctxt map_F_respect_inst THEN' rtac ctxt F_in'D_thm) THEN | |
| 1641 |                   EVERY (map (fn {F_in'_alt2, ...} =>
 | |
| 1642 | unfold_thms_tac ctxt [F_in'_alt2] THEN | |
| 1643 | HEADGOAL (EVERY' | |
| 1644 | [rtac ctxt CollectI, | |
| 1645 | rtac ctxt subset_refl, | |
| 1646 | rtac ctxt ballI, | |
| 1647 | SELECT_GOAL (unfold_thms_tac ctxt [o_apply]), | |
| 1648 |                       rtac ctxt @{thm arg_cong[where f=Inr]},
 | |
| 1649 | asm_full_simp_tac ctxt])) set_F'_thmss) end; | |
| 1650 | in prove lthy (var_x :: var_fs @ var_gs) map_F_cong'_tm map_F_cong'_tac end; | |
| 1651 | ||
| 1652 | (* rel_F'_set: "rel_F' P x y \<longleftrightarrow> | |
| 1653 |             (\<exists>z. set_F' z \<subseteq> {(x, y). P x y} \<and> REL (map_F fst z) x \<and> REL (map_F snd z) y)" *)
 | |
| 1654 | val rel_F'_set_thm = let | |
| 1655 | val lhs = list_comb (#tm rel_F', var_Ps) $ var_x $ var_y; | |
| 1656 |             fun mk_subset_A var_a var_b var_P {set_F', ...} = let
 | |
| 1657 | val collect_A = mk_case_prod (var_a, var_b) (var_P $ var_a $ var_b); | |
| 1658 | in mk_leq (subst_atomic_types (alphas ~~ typ_pairs) (#tm set_F') $ var_z) collect_A end; | |
| 1659 |             val subset_As = @{map 4} mk_subset_A var_as var_bs var_Ps set_F'_aux_defs;
 | |
| 1660 | fun mk_map mfs f z = | |
| 1661 | Term.list_comb (mfs, map (fst o Term.strip_comb o f) var_ts) $ z; | |
| 1662 | val map_F_fst = mk_map_of_bnf deads typ_pairs alphas bnf_F; | |
| 1663 | val map_F_snd = mk_map_of_bnf deads typ_pairs betas bnf_F; | |
| 1664 | val map_fst = equiv_rel_a $ (mk_map map_F_fst HOLogic.mk_fst var_z) $ var_x; | |
| 1665 | val map_snd = equiv_rel_b $ (mk_map map_F_snd HOLogic.mk_snd var_z) $ var_y; | |
| 1666 | val rhs = let val (z, T) = dest_Free var_z in | |
| 1667 | HOLogic.mk_exists (z, T, fold_rev (fn a => fn b => HOLogic.mk_conj (a, b)) | |
| 1668 | (subset_As @ [map_fst]) map_snd) end; | |
| 1669 | val rel_F'_set_tm = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)); | |
| 1670 | ||
| 1671 | val maybePairsTs = map HOLogic.mk_prodT (map mk_MaybeT alphas ~~ map mk_MaybeT betas) | |
| 1672 | fun mk_map_prod_projr aT bT = let | |
| 1673 | val (mabT, (maT, mbT)) = `HOLogic.mk_prodT (apply2 mk_MaybeT (aT, bT)); | |
| 1674 |                 val map_prod_const = Const (@{const_name map_prod},
 | |
| 1675 | (maT --> aT) --> (mbT --> bT) --> mabT --> HOLogic.mk_prodT (aT, bT)); | |
| 1676 | in map_prod_const $ fromJust_const aT $ fromJust_const bT end; | |
| 1677 | ||
| 1678 | fun exI_OF_tac ctxt tm = rtac ctxt | |
| 1679 | (infer_instantiate' ctxt (NONE :: [SOME (Thm.cterm_of ctxt tm)]) exI); | |
| 1680 | ||
| 1681 | (* REL (map_F Inr x) (map_F fst z) \<Longrightarrow> REL (map_F snd z) (map_F Inr y) \<Longrightarrow> | |
| 1682 |                  set_F z \<subseteq> {(x, y). rel_sum (=) P x y} \<Longrightarrow>
 | |
| 1683 |                  \<exists>z. set_F' z \<subseteq> {(x, y). P x y} \<and> REL (map_F fst z) x \<and> REL (map_F snd z) y *)
 | |
| 1684 |             fun subgoal1_tac {context = ctxt, params, ...} =
 | |
| 1685 | let | |
| 1686 | val z = (case params of | |
| 1687 | (_ :: _ :: (_, ct) :: _) => Thm.term_of ct | |
| 1688 | | _ => error "won't happen"); | |
| 1689 | val map_F_projr_z = list_comb (mk_map_of_bnf deads maybePairsTs typ_pairs bnf_F, | |
| 1690 | map2 mk_map_prod_projr alphas betas) $ z; | |
| 1691 | in | |
| 1692 | HEADGOAL (exI_OF_tac ctxt map_F_projr_z) THEN | |
| 1693 |                 HEADGOAL (EVERY' (maps (fn {set_F'_subset, set_F'_respect, set_map_F', ...} =>
 | |
| 1694 | [rtac ctxt conjI, | |
| 1695 |                   dtac ctxt (set_F'_subset RS @{thm order_trans}),
 | |
| 1696 |                   TWICE (dtac ctxt (set_F'_respect RS @{thm rel_funD})),
 | |
| 1697 | SELECT_GOAL (unfold_thms_tac ctxt [set_map_F']), | |
| 1698 |                   etac ctxt @{thm in_rel_sum_in_image_projr},
 | |
| 1699 | TWICE (assume_tac ctxt)]) set_F'_thmss)) THEN | |
| 1700 | HEADGOAL (EVERY' (map (fn Ts => FIRST' | |
| 1701 | [dtac ctxt (map_F_rsp_of (map fromJust_const Ts) ctxt), | |
| 1702 | etac ctxt sym , assume_tac ctxt]) [alphas, betas])) THEN | |
| 1703 | unfold_thms_tac ctxt (map_F_comp :: | |
| 1704 |                   @{thms fst_comp_map_prod snd_comp_map_prod comp_projr_Inr} @ [map_F_id]) THEN
 | |
| 1705 | HEADGOAL (rtac ctxt conjI) THEN | |
| 1706 | HEADGOAL (etac ctxt (#symp qthms) THEN' assume_tac ctxt | |
| 1707 | ORELSE' (EVERY' (maps (fn Ts => | |
| 1708 | [dtac ctxt (map_F_rsp_of (map fromJust_const Ts) ctxt), | |
| 1709 | SELECT_GOAL (unfold_thms_tac ctxt (map_F_comp :: | |
| 1710 |                       @{thms fst_comp_map_prod snd_comp_map_prod comp_projr_Inr} @ [map_F_id])),
 | |
| 1711 | assume_tac ctxt]) [alphas, betas]))) end; | |
| 1712 | ||
| 1713 |             (* set_F' z \<subseteq> {(x, y). P x y} \<Longrightarrow> REL (map_F fst z) x \<Longrightarrow> REL (map_F snd z) y \<Longrightarrow>
 | |
| 1714 | \<exists>b. REL (map_F Inr x) b \<and> (\<exists>ba. rel_F (rel_sum (=) P) b ba \<and> REL ba (map_F Inr y)) *) | |
| 1715 |             fun subgoal2_tac {context = ctxt, params, ...} = let
 | |
| 1716 | val z = (case params of | |
| 1717 | ((_, ct) :: _) => Thm.term_of ct | |
| 1718 | | _ => error "won't happen"); | |
| 1719 | ||
| 1720 | fun exI_map_Ifs_tac mk_proj Ts = exI_OF_tac ctxt (list_comb | |
| 1721 |                   (mk_map_of_bnf deads typ_pairs (map mk_MaybeT Ts) bnf_F, @{map 3}
 | |
| 1722 |                     (fn var_t => fn {set_F', ...} => fn T => lambda var_t (BNF_FP_Util.mk_If
 | |
| 1723 | (HOLogic.mk_mem (var_t, subst_Ts (#tm set_F') typ_pairs $ z)) | |
| 1724 | (mk_Just (mk_proj var_t)) (mk_Nothing T))) var_ts set_F'_aux_defs Ts) $ z) | |
| 1725 | ||
| 1726 | fun mk_REL_trans_map_F n = (rotate_prems n (#transp qthms) OF | |
| 1727 | [rel_funD map_F_respect_thm] OF (replicate live refl @ [#symp qthms])); | |
| 1728 | in | |
| 1729 | HEADGOAL (EVERY' | |
| 1730 | [exI_map_Ifs_tac HOLogic.mk_fst alphas, | |
| 1731 | rtac ctxt conjI, | |
| 1732 | etac ctxt (mk_REL_trans_map_F 0)]) THEN | |
| 1733 |                 unfold_thms_tac ctxt [map_F_comp, @{thm o_def}] THEN
 | |
| 1734 | HEADGOAL (rtac ctxt map_F_cong'_thm) THEN | |
| 1735 |                 REPEAT_DETERM_N live (HEADGOAL (etac ctxt @{thm if_P[symmetric]})) THEN
 | |
| 1736 | HEADGOAL (EVERY' [exI_map_Ifs_tac HOLogic.mk_snd betas, rtac ctxt conjI]) THEN | |
| 1737 | unfold_thms_tac ctxt rel_map THEN | |
| 1738 | HEADGOAL (rtac ctxt rel_refl_strong) THEN | |
| 1739 |                 REPEAT_DETERM_N live (HEADGOAL (etac ctxt @{thm subset_rel_sumI})) THEN
 | |
| 1740 | HEADGOAL (etac ctxt (mk_REL_trans_map_F 1 OF [#symp qthms])) THEN | |
| 1741 |                 unfold_thms_tac ctxt [map_F_comp, @{thm o_def}] THEN
 | |
| 1742 | HEADGOAL (rtac ctxt map_F_cong'_thm) THEN | |
| 1743 |                 REPEAT_DETERM_N live (HEADGOAL (etac ctxt @{thm if_P})) end;
 | |
| 1744 | ||
| 1745 | fun rel_F'_set_tac ctxt = EVERY | |
| 1746 |               ([unfold_thms_tac ctxt (#def rel_F' :: #REL qthms :: @{thms vimage2p_def relcompp_apply}),
 | |
| 1747 | HEADGOAL (rtac ctxt iffI), | |
| 1748 | (HEADGOAL o TWICE) (etac ctxt exE THEN' etac ctxt conjE), | |
| 1749 | HEADGOAL (EVERY' | |
| 1750 | [dtac ctxt (in_rel RS iffD1), | |
| 1751 | etac ctxt exE, | |
| 1752 | TWICE (etac ctxt conjE), | |
| 1753 | etac ctxt CollectE, | |
| 1754 | hyp_subst_tac ctxt]), | |
| 1755 | (REPEAT_DETERM_N (live-1) o HEADGOAL) (etac ctxt conjE), | |
| 1756 | HEADGOAL (Subgoal.FOCUS_PARAMS subgoal1_tac ctxt THEN' etac ctxt exE), | |
| 1757 | (REPEAT_DETERM_N (live+1) o HEADGOAL) (etac ctxt conjE), | |
| 1758 | HEADGOAL (Subgoal.FOCUS_PARAMS subgoal2_tac ctxt)]); | |
| 1759 | ||
| 1760 | in prove lthy (var_x :: var_y :: var_Ps) rel_F'_set_tm rel_F'_set_tac end; | |
| 1761 | ||
| 1762 | (* tactics *) | |
| 1763 | ||
| 1764 | (* map_G_id0: abs_G \<circ> map_F id \<circ> rep_G = id *) | |
| 1765 | fun map_G_id0_tac ctxt = HEADGOAL (EVERY' [SELECT_GOAL (unfold_thms_tac ctxt | |
| 1766 |             [@{thm fun_eq_iff}, o_apply, map_F_id0, id_apply, #abs_rep qthms]),
 | |
| 1767 | rtac ctxt allI, rtac ctxt refl]); | |
| 1768 | ||
| 1769 | (* map_G (g \<circ> f) = map_G g \<circ> map_G f *) | |
| 1770 | fun map_G_comp0_tac ctxt = HEADGOAL (EVERY' [rtac ctxt ext, rtac ctxt sym, | |
| 1771 | SELECT_GOAL (unfold_thms_tac ctxt [o_apply, map_F_comp0]), rtac ctxt (#rel_abs qthms), | |
| 1772 | rtac ctxt map_F_rsp, rtac ctxt (#rep_abs qthms), rtac ctxt (#reflp qthms)]); | |
| 1773 | ||
| 1774 | (* map_G_cong: (\<And>z. z \<in> set_G x \<Longrightarrow> f z = g z) \<Longrightarrow> map_G f x = map_G g x *) | |
| 1775 | fun map_G_cong_tac ctxt = EVERY | |
| 1776 | [Local_Defs.fold_tac ctxt (map #set_F'_def set_F'_thmss), | |
| 1777 | unfold_thms_tac ctxt [o_apply], | |
| 1778 | HEADGOAL (rtac ctxt (#rel_abs qthms) THEN' rtac ctxt map_F_cong'_thm), | |
| 1779 | REPEAT_DETERM_N live (HEADGOAL (asm_full_simp_tac ctxt))]; | |
| 1780 | ||
| 1781 | (* set_G_map0_G: set_G \<circ> map_G f = f ` set_G *) | |
| 1782 | fun mk_set_G_map0_G_tac thms ctxt = | |
| 1783 | HEADGOAL (rtac ctxt ext) THEN | |
| 1784 | EVERY [unfold_thms_tac ctxt [o_apply], | |
| 1785 | Local_Defs.fold_tac ctxt [#set_F'_def thms]] THEN | |
| 1786 | HEADGOAL (EVERY' (map (rtac ctxt) | |
| 1787 | [trans OF [#set_map_F' thms RS sym, sym] RS sym, | |
| 1788 |                @{thm rel_funD} OF [#set_F'_respect thms],
 | |
| 1789 | #rep_abs qthms, | |
| 1790 | map_F_rsp, | |
| 1791 | #rep_reflp qthms])); | |
| 1792 | ||
| 1793 | (* bd_card_order: card_order bd_F *) | |
| 1794 | fun bd_card_order_tac ctxt = HEADGOAL (rtac ctxt bd_card_order); | |
| 1795 | ||
| 1796 | (* bd_cinfinite: BNF_Cardinal_Arithmetic.cinfinite bd_F *) | |
| 1797 | fun bd_cinfinite_tac ctxt = HEADGOAL (rtac ctxt bd_cinfinite); | |
| 1798 | ||
| 75624 | 1799 | (* bd_regularCard: regularCard bd_F *) | 
| 1800 | fun bd_regularCard_tac ctxt = HEADGOAL (rtac ctxt bd_regularCard); | |
| 1801 | ||
| 1802 | (*target: ordLess2 (card_of (set_F' (rep_G x_))) bd_F*) | |
| 71262 | 1803 | fun mk_set_G_bd_tac thms set_bd_thm ctxt = EVERY | 
| 1804 | [Local_Defs.fold_tac ctxt [#set_F'_def thms], | |
| 1805 | unfold_thms_tac ctxt [o_apply], | |
| 75624 | 1806 |             HEADGOAL (rtac ctxt (@{thm ordLeq_ordLess_trans} OF
 | 
| 71262 | 1807 |               [@{thm card_of_mono1} OF [#set_F'_subset thms], set_bd_thm]))];
 | 
| 1808 | ||
| 1809 | (* rel_compp: rel_G R OO rel_G S \<le> rel_G (R OO S) *) | |
| 1810 | fun rel_compp_tac ctxt = EVERY | |
| 1811 | [unfold_thms_tac ctxt [#REL qthms], | |
| 1812 |             HEADGOAL (TWICE (rtac ctxt @{thm vimage2p_relcompp_mono})),
 | |
| 1813 |             (unfold_thms_tac ctxt (REL_OO_REL_left_thm :: @{thms relcompp_assoc})),
 | |
| 1814 |             (unfold_thms_tac ctxt [Local_Defs.unfold ctxt @{thms eq_OO}
 | |
| 1815 | (infer_instantiate' ctxt [HOLogic.eq_const HOLogic.unitT |> Thm.cterm_of ctxt |> SOME] | |
| 1816 |                 @{thm sum.rel_compp})]),
 | |
| 1817 | HEADGOAL (rtac ctxt rel_pos_distr_thm), | |
| 1818 | unfold_thms_tac ctxt | |
| 1819 |               @{thms fun_eq_iff bot_apply bot_bool_def not_all eq_False not_not OO_def},
 | |
| 1820 |             REPEAT_DETERM (HEADGOAL (resolve_tac ctxt [exI, conjI, @{thm rel_sum.intros(1)}, refl]))];
 | |
| 1821 | ||
| 1822 |           (* rel_G R_ = (\<lambda>x y. \<exists>z. set_G z \<subseteq> {(x, y). R x y} \<and> map_G fst z = x \<and> map_G snd z = y) *)
 | |
| 1823 | fun rel_compp_Grp_tac ctxt = let | |
| 1824 | val _ = () | |
| 1825 |             in EVERY [Local_Defs.fold_tac ctxt (@{thm Grp_def} :: map #set_F'_def set_F'_thmss),
 | |
| 1826 | unfold_thms_tac ctxt | |
| 1827 |                 [o_apply, @{thm mem_Collect_eq}, @{thm OO_Grp_alt}, @{thm vimage2p_def}],
 | |
| 1828 |               Local_Defs.fold_tac ctxt [Local_Defs.unfold ctxt @{thms vimage2p_def} (#def rel_F')],
 | |
| 1829 | unfold_thms_tac ctxt [rel_F'_set_thm], | |
| 1830 | HEADGOAL (TWICE (rtac ctxt ext)), | |
| 1831 | HEADGOAL (rtac ctxt iffI), | |
| 1832 | REPEAT_DETERM (ALLGOALS (eresolve_tac ctxt [exE, conjE])), | |
| 1833 | HEADGOAL (rtac ctxt exI), | |
| 1834 | REPEAT_FIRST (resolve_tac ctxt [conjI]), | |
| 1835 |               HEADGOAL (EVERY' (maps (fn {set_F'_respect, ...} =>
 | |
| 1836 |                 [etac ctxt @{thm subset_trans[rotated]},
 | |
| 1837 | rtac ctxt equalityD1, | |
| 1838 |                 rtac ctxt (@{thm rel_funD} OF [set_F'_respect]),
 | |
| 1839 | rtac ctxt (#rep_abs qthms), | |
| 1840 | rtac ctxt (#reflp qthms)]) set_F'_thmss)), | |
| 1841 | (HEADGOAL o TWICE o EVERY') | |
| 1842 | [rtac ctxt (trans OF [asm_rl, #abs_rep qthms]), | |
| 1843 | rtac ctxt (#rel_abs qthms), | |
| 1844 | etac ctxt (rotate_prems 1 (#transp qthms)), | |
| 1845 | rtac ctxt map_F_rsp, | |
| 1846 | rtac ctxt (#rep_abs qthms), | |
| 1847 | rtac ctxt (#reflp qthms) | |
| 1848 | ], | |
| 1849 | HEADGOAL (rtac ctxt exI THEN' rtac ctxt conjI), | |
| 1850 | (REPEAT_DETERM_N live o HEADGOAL o EVERY') | |
| 1851 | [assume_tac ctxt, rtac ctxt conjI], | |
| 1852 | (HEADGOAL o TWICE o EVERY') [ | |
| 1853 | hyp_subst_tac ctxt, | |
| 1854 | rtac ctxt (#rep_abs_rsp qthms), | |
| 1855 | rtac ctxt map_F_rsp, | |
| 1856 | rtac ctxt (#rep_reflp qthms)]] | |
| 1857 | end; | |
| 1858 | ||
| 1859 | fun pred_G_set_G_tac ctxt = HEADGOAL (rtac ctxt refl); | |
| 1860 | ||
| 1861 | val tactics = map_G_id0_tac :: map_G_comp0_tac :: map_G_cong_tac :: | |
| 1862 | map mk_set_G_map0_G_tac set_F'_thmss @ | |
| 75624 | 1863 | bd_card_order_tac :: bd_cinfinite_tac :: bd_regularCard_tac :: | 
| 71262 | 1864 | map2 mk_set_G_bd_tac set_F'_thmss set_bd_thms @ | 
| 1865 | rel_compp_tac :: rel_compp_Grp_tac :: [pred_G_set_G_tac]; | |
| 1866 | ||
| 71469 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 1867 | (* val wits_G = [abs_G o wit_F] *) | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 1868 | val (wits_G, wit_thms) = | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 1869 | let | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 1870 | val wit_F_thmss = map (map2 (fn set_F' => fn wit => | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 1871 | (#set_F'_subset set_F' RS set_mp RS wit) | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 1872 | |> unfold_thms lthy [#set_F'_def set_F']) set_F'_thmss) | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 1873 | (wit_thmss_of_bnf bnf_F); | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 1874 | val (wits_F, wit_F_thmss) = split_list | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 1875 | (filter_out (fn ((J, _), _) => exists (fn (I, _) => subset (op =) (I, J)) Iwits) | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 1876 | (wits_F ~~ wit_F_thmss)); | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 1877 | fun mk_wit (I, wit) = let val vars = (map (fn n => nth var_as n) I) | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 1878 | in fold_rev lambda vars (abs_G $ list_comb (wit, vars)) end; | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 1879 | in | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 1880 | (map mk_wit (Iwits @ wits_F), wit_thmss @ flat wit_F_thmss) | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 1881 | end; | 
| 
d7ef73df3d15
lift BNF witnesses for quotients (unless better ones are specified by the user)
 traytel parents: 
71393diff
changeset | 1882 | |
| 71262 | 1883 |           fun mk_wit_tacs ({set_F'_def, set_F'_respect, ...} :: set_F'_thmss) (w :: ws) ctxt =
 | 
| 1884 |                 EVERY [unfold_thms_tac ctxt [@{thm o_def},
 | |
| 1885 |                     set_F'_respect RS @{thm rel_funD} OF [#rep_abs qthms OF [(#reflp qthms)]]],
 | |
| 1886 | unfold_thms_tac ctxt [set_F'_def], | |
| 1887 | HEADGOAL (etac ctxt w)] | |
| 1888 | THEN mk_wit_tacs set_F'_thmss ws ctxt | |
| 1889 | | mk_wit_tacs [] ws ctxt = mk_wit_tacs set_F'_thmss ws ctxt | |
| 1890 | | mk_wit_tacs _ _ _ = all_tac; | |
| 1891 | ||
| 1892 | val (bnf_G, lthy) = bnf_def Dont_Inline (user_policy Note_Some) true I | |
| 1893 | tactics (mk_wit_tacs [] wit_thms) NONE map_b rel_b pred_b set_bs | |
| 1894 | (((((((Binding.empty, absT), map_G), sets_G), bd_G), wits_G), SOME rel_G), NONE) lthy; | |
| 1895 | ||
| 1896 | val old_defs = | |
| 1897 |             {sets = set_defs_of_bnf bnf_G, map = map_def_of_bnf bnf_G, rel = rel_def_of_bnf bnf_G};
 | |
| 1898 | ||
| 1899 | val set_F'_defs = map (mk_abs_def o #set_F'_def) set_F'_thmss; | |
| 1900 | val unfold_morphism = Morphism.thm_morphism "BNF" | |
| 1901 | (unfold_thms lthy (defs @ #def REL :: set_F'_defs)); | |
| 1902 | val (bnf_G, lthy) = morph_bnf_defs unfold_morphism bnf_G | |
| 1903 | |> (fn bnf => note_bnf_defs bnf lthy); | |
| 1904 | ||
| 1905 | (* auxiliary lemmas transfer for transfer *) | |
| 1906 |           val rel_monoD_rotated = rotate_prems ~1 (rel_mono_of_bnf bnf_F RS @{thm predicate2D});
 | |
| 1907 | ||
| 1908 | val REL_pos_distrI = let | |
| 1909 | fun tac ctxt = EVERY | |
| 1910 |                 [HEADGOAL (dtac ctxt (rotate_prems ~1 (rel_pos_distr_thm RS @{thm predicate2D}))),
 | |
| 1911 | (REPEAT_DETERM o HEADGOAL) (rtac ctxt conjI ORELSE' assume_tac ctxt), | |
| 1912 |                 (REPEAT_DETERM o HEADGOAL) (etac ctxt @{thm relcomppE}),
 | |
| 1913 | HEADGOAL (dtac ctxt rel_monoD_rotated), | |
| 1914 | (REPEAT_DETERM o HEADGOAL) | |
| 1915 |                   (assume_tac ctxt ORELSE' rtac ctxt @{thm relcomppI})];
 | |
| 1916 | in prove lthy (var_x :: var_y' :: var_Ps @ var_Qs @ var_Rs) REL_pos_distrI_tm tac end; | |
| 1917 | ||
| 1918 | val rel_F_rel_F' = let | |
| 1919 | val rel_F = mk_rel_of_bnf deads alphas betas bnf_F; | |
| 1920 | val rel_F_rel_F'_tm = (rel_F, #tm rel_F') | |
| 1921 | |> apply2 (fn R => HOLogic.mk_Trueprop (list_comb (R, var_Ps) $ var_x $ var_y)) | |
| 1922 | |> Logic.mk_implies; | |
| 1923 | fun rel_F_rel_F'_tac ctxt = EVERY | |
| 1924 | [HEADGOAL (dtac ctxt (in_rel_of_bnf bnf_F RS iffD1)), | |
| 1925 |                 unfold_thms_tac ctxt (rel_F'_set_thm :: @{thms mem_Collect_eq}),
 | |
| 1926 | (REPEAT_DETERM o HEADGOAL) (eresolve_tac ctxt [exE, conjE]), | |
| 1927 | HEADGOAL (rtac ctxt exI), | |
| 1928 | HEADGOAL (EVERY' (maps (fn thms => | |
| 1929 | [rtac ctxt conjI, | |
| 1930 | rtac ctxt subsetI, | |
| 1931 | dtac ctxt (set_mp OF [#set_F'_subset thms]), | |
| 1932 | dtac ctxt subsetD, | |
| 1933 | assume_tac ctxt, assume_tac ctxt]) set_F'_thmss)), | |
| 1934 | (REPEAT_DETERM o HEADGOAL) | |
| 1935 | (rtac ctxt conjI ORELSE' hyp_subst_tac ctxt THEN' rtac ctxt (#reflp qthms))] | |
| 1936 | in prove lthy (var_x :: var_y :: var_Ps) rel_F_rel_F'_tm rel_F_rel_F'_tac end; | |
| 1937 | ||
| 1938 | fun inst_REL_pos_distrI n vs aTs bTs ctxt = | |
| 1939 | infer_instantiate' ctxt (replicate n NONE @ (rel_Maybes vs aTs bTs | |
| 1940 | |> map (SOME o Thm.cterm_of ctxt))) REL_pos_distrI; | |
| 1941 | ||
| 1942 |           val Tss = {abs = typ_subst_atomic (alphas ~~ betas) absT, rep = repT, Ds0 = map TFree Ds0,
 | |
| 1943 | deads = deads, alphas = alphas, betas = betas, gammas = gammas, deltas = deltas}; | |
| 1944 | ||
| 1945 | val thms = | |
| 1946 |             {map_F_rsp = map_F_rsp,
 | |
| 1947 | rel_F'_def = #def rel_F', | |
| 1948 | rel_F_rel_F' = rel_F_rel_F', | |
| 1949 | rel_F'_set = rel_F'_set_thm, | |
| 1950 | rel_monoD_rotated = rel_monoD_rotated} | |
| 1951 | ||
| 1952 | val transfer_consts = mk_quotient_transfer_tacs bnf_F Tss live | |
| 1953 | qthms thms set_F'_thmss old_defs inst_REL_pos_distrI | |
| 1954 | map_raw rel_raw (map (#tm o #set_F') set_F'_aux_defs); | |
| 1955 | val quiet = exists (fn No_Warn_Transfer => true | _ => false) opts; | |
| 1956 | in | |
| 1957 | lthy |> BNF_Def.register_bnf plugins absT_name bnf_G |> | |
| 1958 | mk_transfer_thms quiet bnf_F bnf_G absT_name transfer_consts (Quotient equiv_thm) Tss | |
| 1959 | (defs @ #def REL :: set_F'_defs) | |
| 1960 | end | |
| 1961 | | _ => raise Match); | |
| 1962 | ||
| 1963 | in (goals, after_qed, #def REL :: defs, lthy) end; | |
| 1964 | ||
| 1965 | ||
| 1966 | (** main commands **) | |
| 61067 | 1967 | |
| 1968 | local | |
| 1969 | ||
| 60918 | 1970 | fun prepare_common prepare_name prepare_sort prepare_term prepare_thm | 
| 71494 | 1971 | (((((plugins, raw_specs), raw_absT_name), raw_wits), xthms_opt), (map_b, rel_b, pred_b)) lthy = | 
| 60918 | 1972 | let | 
| 71262 | 1973 | val absT_name = prepare_name lthy raw_absT_name; | 
| 1974 | ||
| 71494 | 1975 | fun bad_input input = | 
| 1976 |       Pretty.chunks [Pretty.para ("Expected theorem(s) of either form:"),
 | |
| 1977 |       Pretty.item [Pretty.enum "," "" "" [Syntax.pretty_term lthy @{term "Quotient R Abs Rep T"},
 | |
| 1978 |         Syntax.pretty_term lthy @{term "reflp R"}]],
 | |
| 1979 |       Pretty.item [Syntax.pretty_term lthy @{term "Quotient R Abs Rep T"}],
 | |
| 1980 |       Pretty.item [Syntax.pretty_term lthy @{term "type_definition Rep Abs A"}],
 | |
| 1981 |       Pretty.para ("Got"), Pretty.enum "," "" "" (map (Thm.pretty_thm lthy) input)]
 | |
| 1982 | |> Pretty.string_of | |
| 1983 | |> error; | |
| 1984 | ||
| 1985 | fun no_refl qthm = | |
| 1986 |       Pretty.chunks [Pretty.para ("Could not find a reflexivity rule for the Quotient theorem:"),
 | |
| 1987 | Pretty.item [Thm.pretty_thm lthy qthm], | |
| 1988 |       Pretty.para ("Try supplying a reflexivity theorem manually or registering it in setup_lifting.")]
 | |
| 1989 | |> Pretty.string_of | |
| 1990 | |> error; | |
| 1991 | ||
| 1992 | fun find_equiv_thm_via_Quotient qthm = | |
| 1993 | let | |
| 1994 | val refl_thms = Lifting_Info.get_reflexivity_rules lthy | |
| 1995 |          |> map (unfold_thms lthy @{thms reflp_eq[symmetric]});
 | |
| 1996 | in | |
| 1997 |         (case refl_thms RL [qthm RS @{thm Quotient_reflp_imp_equivp}] of
 | |
| 1998 | [] => no_refl qthm | |
| 1999 | | thm :: _ => thm) | |
| 2000 | end; | |
| 2001 | ||
| 2002 | val (lift_thm, equiv_thm) = | |
| 2003 | (case Option.map (prepare_thm lthy) xthms_opt of | |
| 2004 | SOME (thms as [qthm, _]) => | |
| 2005 |           (case try (fn thms => @{thm Quotient_reflp_imp_equivp} OF thms) thms of
 | |
| 2006 |             SOME equiv_thm => (qthm RS @{thm Quotient_Quotient3}, Quotient equiv_thm)
 | |
| 2007 | | NONE => bad_input thms) | |
| 2008 | | SOME [thm] => | |
| 2009 |           (case try (fn thm => thm RS @{thm Quotient_Quotient3}) thm of
 | |
| 2010 | SOME qthm => (qthm, Quotient (find_equiv_thm_via_Quotient thm)) | |
| 2011 |           | NONE => if can (fn thm => thm RS @{thm type_definition.Rep}) thm
 | |
| 2012 | then (thm, Typedef) | |
| 2013 | else bad_input [thm]) | |
| 2014 | | NONE => (case Lifting_Info.lookup_quotients lthy absT_name of | |
| 2015 |             SOME {quot_thm = qthm, ...} =>
 | |
| 71558 
1cf958713cf7
remove Thm.transfer workaround made obsplete by cf2406e654cf
 traytel parents: 
71494diff
changeset | 2016 |               (case [qthm] RL @{thms Quotient_eq_onp_typedef Quotient_eq_onp_type_copy} of
 | 
| 
1cf958713cf7
remove Thm.transfer workaround made obsplete by cf2406e654cf
 traytel parents: 
71494diff
changeset | 2017 | thm :: _ => (thm, Typedef) | 
| 
1cf958713cf7
remove Thm.transfer workaround made obsplete by cf2406e654cf
 traytel parents: 
71494diff
changeset | 2018 |               | _ => (qthm RS @{thm Quotient_Quotient3},
 | 
| 
1cf958713cf7
remove Thm.transfer workaround made obsplete by cf2406e654cf
 traytel parents: 
71494diff
changeset | 2019 | Quotient (find_equiv_thm_via_Quotient qthm))) | 
| 71494 | 2020 | | NONE => (Typedef.get_info lthy absT_name |> hd |> snd |> #type_definition, Typedef)) | 
| 2021 | | SOME thms => bad_input thms); | |
| 61067 | 2022 | val wits = (Option.map o map) (prepare_term lthy) raw_wits; | 
| 2023 | val specs = | |
| 71262 | 2024 |       map (apsnd (apsnd (the_default @{sort type} o Option.map (prepare_sort lthy)))) raw_specs;
 | 
| 60918 | 2025 | |
| 71494 | 2026 | val which_bnf = (case equiv_thm of | 
| 2027 | Quotient thm => quotient_bnf (thm, lift_thm) | |
| 2028 | | Typedef => typedef_bnf lift_thm); | |
| 60918 | 2029 | in | 
| 71494 | 2030 | which_bnf wits specs map_b rel_b pred_b plugins lthy | 
| 60918 | 2031 | end; | 
| 2032 | ||
| 2033 | fun prepare_lift_bnf prepare_name prepare_sort prepare_term prepare_thm = | |
| 2034 | (fn (goals, after_qed, definitions, lthy) => | |
| 2035 | lthy | |
| 2036 | |> Proof.theorem NONE after_qed (map (single o rpair []) goals) | |
| 61841 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
61073diff
changeset | 2037 | |> Proof.refine_singleton | 
| 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
61073diff
changeset | 2038 | (Method.Basic (fn ctxt => SIMPLE_METHOD (unfold_thms_tac ctxt definitions))) | 
| 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
61073diff
changeset | 2039 | |> Proof.refine_singleton (Method.primitive_text (K I))) oo | 
| 60918 | 2040 | prepare_common prepare_name prepare_sort prepare_term prepare_thm o apfst (apfst (apsnd SOME)); | 
| 2041 | ||
| 2042 | fun prepare_solve prepare_name prepare_typ prepare_sort prepare_thm tacs = | |
| 63023 
1f4b011c5738
unfold internal definitions before emitting a proof obligation
 traytel parents: 
62969diff
changeset | 2043 | (fn (goals, after_qed, definitions, lthy) => | 
| 60918 | 2044 | lthy | 
| 71262 | 2045 | |> after_qed (map2 (fn goal => fn tac => [Goal.prove_sorry lthy [] [] goal | 
| 63023 
1f4b011c5738
unfold internal definitions before emitting a proof obligation
 traytel parents: 
62969diff
changeset | 2046 |         (fn (ctxtprems as {context = ctxt, prems = _}) =>
 | 
| 
1f4b011c5738
unfold internal definitions before emitting a proof obligation
 traytel parents: 
62969diff
changeset | 2047 | unfold_thms_tac ctxt definitions THEN tac ctxtprems)]) | 
| 
1f4b011c5738
unfold internal definitions before emitting a proof obligation
 traytel parents: 
62969diff
changeset | 2048 | goals (tacs (length goals)))) oo | 
| 62777 | 2049 | prepare_common prepare_name prepare_typ prepare_sort prepare_thm; | 
| 60918 | 2050 | |
| 61067 | 2051 | in | 
| 2052 | ||
| 2053 | val lift_bnf_cmd = | |
| 2054 | prepare_lift_bnf | |
| 80634 
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
 wenzelm parents: 
79564diff
changeset | 2055 |     (dest_Type_name oo Proof_Context.read_type_name {proper = true, strict = false})
 | 
| 71494 | 2056 | Syntax.read_sort Syntax.read_term Attrib.eval_thms; | 
| 61067 | 2057 | |
| 62777 | 2058 | fun lift_bnf args tacs = | 
| 2059 | prepare_solve (K I) (K I) (K I) (K I) (K tacs) args; | |
| 61067 | 2060 | |
| 71262 | 2061 | fun copy_bnf_tac {context = ctxt, prems = _} =
 | 
| 2062 | REPEAT_DETERM (resolve_tac ctxt [bexI, conjI, UNIV_I, refl, subset_refl] 1); | |
| 2063 | ||
| 61067 | 2064 | val copy_bnf = | 
| 62777 | 2065 | apfst (apfst (rpair NONE)) | 
| 71494 | 2066 | #> apfst (apsnd (Option.map single)) | 
| 62777 | 2067 | #> prepare_solve (K I) (K I) (K I) (K I) | 
| 71262 | 2068 | (fn n => replicate n copy_bnf_tac); | 
| 61067 | 2069 | |
| 2070 | val copy_bnf_cmd = | |
| 62777 | 2071 | apfst (apfst (rpair NONE)) | 
| 71494 | 2072 | #> apfst (apsnd (Option.map single)) | 
| 62777 | 2073 | #> prepare_solve | 
| 80634 
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
 wenzelm parents: 
79564diff
changeset | 2074 |     (dest_Type_name oo Proof_Context.read_type_name {proper = true, strict = false})
 | 
| 71494 | 2075 | Syntax.read_sort Syntax.read_term Attrib.eval_thms | 
| 71262 | 2076 | (fn n => replicate n copy_bnf_tac); | 
| 61067 | 2077 | |
| 2078 | end; | |
| 2079 | ||
| 71262 | 2080 | (** outer syntax **) | 
| 61067 | 2081 | |
| 2082 | local | |
| 60918 | 2083 | |
| 71262 | 2084 | (* parsers *) | 
| 2085 | ||
| 60918 | 2086 | val parse_wits = | 
| 71262 | 2087 |   @{keyword "["} |-- (Parse.name --| @{keyword ":"} -- Scan.repeat Parse.term >>
 | 
| 60918 | 2088 |     (fn ("wits", Ts) => Ts
 | 
| 2089 |       | (s, _) => error ("Unknown label " ^ quote s ^ " (expected \"wits\")"))) --|
 | |
| 71262 | 2090 |   @{keyword "]"} || Scan.succeed [];
 | 
| 60918 | 2091 | |
| 71262 | 2092 | fun parse_common_opts p = | 
| 2093 |   Scan.optional (@{keyword "("} |--
 | |
| 60918 | 2094 | Parse.list1 (Parse.group (K "option") | 
| 71262 | 2095 | (Scan.first (p :: [Plugin_Name.parse_filter >> Plugins_Option, | 
| 2096 | Parse.reserved "no_warn_transfer" >> K No_Warn_Transfer]))) | |
| 2097 |     --| @{keyword ")"}) [];
 | |
| 60918 | 2098 | |
| 71262 | 2099 | val parse_lift_opts = Parse.reserved "no_warn_wits" >> K No_Warn_Wits |> parse_common_opts; | 
| 60918 | 2100 | |
| 71262 | 2101 | val parse_copy_opts = parse_common_opts Scan.fail; | 
| 2102 | ||
| 2103 | val parse_xthm = Scan.option (Parse.reserved "via" |-- Parse.thm); | |
| 71494 | 2104 | val parse_xthms = Scan.option (Parse.reserved "via" |-- Parse.thms1); | 
| 60918 | 2105 | |
| 61067 | 2106 | in | 
| 2107 | ||
| 71262 | 2108 | (* exposed commands *) | 
| 60918 | 2109 | |
| 2110 | val _ = | |
| 71262 | 2111 |   Outer_Syntax.local_theory_to_proof @{command_keyword lift_bnf}
 | 
| 2112 | "register a quotient type/subtype of a bounded natural functor (BNF) as a BNF" | |
| 2113 | ((parse_lift_opts -- parse_type_args_named_constrained -- Parse.type_const -- parse_wits -- | |
| 71494 | 2114 | parse_xthms -- parse_map_rel_pred_bindings) >> lift_bnf_cmd); | 
| 71262 | 2115 | |
| 2116 | val _ = | |
| 2117 |   Outer_Syntax.local_theory @{command_keyword copy_bnf}
 | |
| 60918 | 2118 | "register a type copy of a bounded natural functor (BNF) as a BNF" | 
| 71262 | 2119 | ((parse_copy_opts -- parse_type_args_named_constrained -- Parse.type_const -- | 
| 2120 | parse_xthm -- parse_map_rel_pred_bindings) >> copy_bnf_cmd); | |
| 60918 | 2121 | |
| 61067 | 2122 | end; | 
| 2123 | ||
| 2124 | end; |