src/HOL/Tools/Transfer/transfer_bnf.ML
changeset 59621 291934bac95e
parent 59582 0fbed69ff081
child 59823 a03696dc3283
equal deleted inserted replaced
59620:92d7d8e4f1bf 59621:291934bac95e
   342     val old_lthy = lthy
   342     val old_lthy = lthy
   343     val (As, lthy) = mk_TFrees' (sorts_of_fp fp_sugar) lthy
   343     val (As, lthy) = mk_TFrees' (sorts_of_fp fp_sugar) lthy
   344     val predTs = map mk_pred1T As
   344     val predTs = map mk_pred1T As
   345     val (preds, lthy) = mk_Frees "P" predTs lthy
   345     val (preds, lthy) = mk_Frees "P" predTs lthy
   346     val args = map mk_eq_onp preds
   346     val args = map mk_eq_onp preds
   347     val cTs = map (SOME o Proof_Context.ctyp_of lthy) (maps (replicate 2) As)
   347     val cTs = map (SOME o Thm.ctyp_of lthy) (maps (replicate 2) As)
   348     val cts = map (SOME o Proof_Context.cterm_of lthy) args
   348     val cts = map (SOME o Thm.cterm_of lthy) args
   349     fun get_rhs thm = thm |> Thm.concl_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd
   349     fun get_rhs thm = thm |> Thm.concl_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd
   350     fun is_eqn thm = can get_rhs thm
   350     fun is_eqn thm = can get_rhs thm
   351     fun rel2pred_massage thm =
   351     fun rel2pred_massage thm =
   352       let
   352       let
   353         val live_step = @{lemma "x = y \<Longrightarrow> (eq_onp P a a \<and> x) = (P a \<and> y)" by (simp only: eq_onp_same_args)}
   353         val live_step = @{lemma "x = y \<Longrightarrow> (eq_onp P a a \<and> x) = (P a \<and> y)" by (simp only: eq_onp_same_args)}