src/HOL/HOLCF/Cfun.thy
changeset 57954 c52223cd1003
parent 57945 cacb00a569e0
child 58880 0baae4311a9f
equal deleted inserted replaced
57953:69728243a614 57954:c52223cd1003
   146       val f = (snd o dest o snd o dest) ct;
   146       val f = (snd o dest o snd o dest) ct;
   147       val [T, U] = Thm.dest_ctyp (ctyp_of_term f);
   147       val [T, U] = Thm.dest_ctyp (ctyp_of_term f);
   148       val tr = instantiate' [SOME T, SOME U] [SOME f]
   148       val tr = instantiate' [SOME T, SOME U] [SOME f]
   149           (mk_meta_eq @{thm Abs_cfun_inverse2});
   149           (mk_meta_eq @{thm Abs_cfun_inverse2});
   150       val rules = Named_Theorems.get ctxt @{named_theorems cont2cont};
   150       val rules = Named_Theorems.get ctxt @{named_theorems cont2cont};
   151       val tac = SOLVED' (REPEAT_ALL_NEW (match_tac rules));
   151       val tac = SOLVED' (REPEAT_ALL_NEW (match_tac (rev rules)));
   152     in SOME (perhaps (SINGLE (tac 1)) tr) end
   152     in SOME (perhaps (SINGLE (tac 1)) tr) end
   153 *}
   153 *}
   154 
   154 
   155 text {* Eta-equality for continuous functions *}
   155 text {* Eta-equality for continuous functions *}
   156 
   156