# HG changeset patch # User nipkow # Date 1476723225 -7200 # Node ID 2e94501cbc347bfe49ed799ae1bccddc229dd595 # Parent 912573107a2e9a2f51871e55b58c61085aa56f77# Parent f76b6dda2e569fefa0a5117ccf34f88fe2ae16a5 merged diff -r f76b6dda2e56 -r 2e94501cbc34 src/HOL/Nonstandard_Analysis/StarDef.thy --- a/src/HOL/Nonstandard_Analysis/StarDef.thy Mon Oct 17 17:33:07 2016 +0200 +++ b/src/HOL/Nonstandard_Analysis/StarDef.thy Mon Oct 17 18:53:45 2016 +0200 @@ -78,6 +78,28 @@ "P \ eventually (\n. Q) \ \ Trueprop P \ Trueprop Q" by (simp add: FreeUltrafilterNat.proper) +text \Standard principles that play a central role in the transfer tactic.\ +definition + Ifun :: "('a \ 'b) star \ 'a star \ 'b star" ("_ \ _" [300,301] 300) where + "Ifun f \ \x. Abs_star + (\F\Rep_star f. \X\Rep_star x. starrel``{\n. F n (X n)})" + +lemma Ifun_congruent2: + "congruent2 starrel starrel (\F X. starrel``{\n. F n (X n)})" +by (auto simp add: congruent2_def equiv_starrel_iff elim!: eventually_rev_mp) + +lemma Ifun_star_n: "star_n F \ star_n X = star_n (\n. F n (X n))" +by (simp add: Ifun_def star_n_def Abs_star_inverse starrel_in_star + UN_equiv_class2 [OF equiv_starrel equiv_starrel Ifun_congruent2]) + +lemma transfer_Ifun: + "\f \ star_n F; x \ star_n X\ \ f \ x \ star_n (\n. F n (X n))" +by (simp only: Ifun_star_n) + +definition + star_of :: "'a \ 'a star" where + "star_of x == star_n (\n. x)" + text \Initialize transfer tactic.\ ML_file "transfer.ML" @@ -155,47 +177,23 @@ subsection \Standard elements\ definition - star_of :: "'a \ 'a star" where - "star_of x == star_n (\n. x)" - -definition Standard :: "'a star set" where "Standard = range star_of" text \Transfer tactic should remove occurrences of @{term star_of}\ setup \Transfer_Principle.add_const @{const_name star_of}\ -declare star_of_def [transfer_intro] - lemma star_of_inject: "(star_of x = star_of y) = (x = y)" by (transfer, rule refl) lemma Standard_star_of [simp]: "star_of x \ Standard" by (simp add: Standard_def) - subsection \Internal functions\ -definition - Ifun :: "('a \ 'b) star \ 'a star \ 'b star" ("_ \ _" [300,301] 300) where - "Ifun f \ \x. Abs_star - (\F\Rep_star f. \X\Rep_star x. starrel``{\n. F n (X n)})" - -lemma Ifun_congruent2: - "congruent2 starrel starrel (\F X. starrel``{\n. F n (X n)})" -by (auto simp add: congruent2_def equiv_starrel_iff elim!: eventually_rev_mp) - -lemma Ifun_star_n: "star_n F \ star_n X = star_n (\n. F n (X n))" -by (simp add: Ifun_def star_n_def Abs_star_inverse starrel_in_star - UN_equiv_class2 [OF equiv_starrel equiv_starrel Ifun_congruent2]) - text \Transfer tactic should remove occurrences of @{term Ifun}\ setup \Transfer_Principle.add_const @{const_name Ifun}\ -lemma transfer_Ifun [transfer_intro]: - "\f \ star_n F; x \ star_n X\ \ f \ x \ star_n (\n. F n (X n))" -by (simp only: Ifun_star_n) - lemma Ifun_star_of [simp]: "star_of f \ star_of x = star_of (f x)" by (transfer, rule refl) diff -r f76b6dda2e56 -r 2e94501cbc34 src/HOL/Nonstandard_Analysis/transfer.ML --- a/src/HOL/Nonstandard_Analysis/transfer.ML Mon Oct 17 17:33:07 2016 +0200 +++ b/src/HOL/Nonstandard_Analysis/transfer.ML Mon Oct 17 18:53:45 2016 +0200 @@ -50,6 +50,25 @@ unstar term end +local exception MATCH +in +fun transfer_star_tac ctxt = + let + fun thm_of (Const (@{const_name Ifun}, _) $ t $ u) = @{thm transfer_Ifun} OF [thm_of t, thm_of u] + | thm_of (Const (@{const_name star_of}, _) $ _) = @{thm star_of_def} + | thm_of (Const (@{const_name star_n}, _) $ _) = @{thm Pure.reflexive} + | thm_of _ = raise MATCH; + + fun thm_of_goal (Const (@{const_name Pure.eq}, _) $ t $ (Const (@{const_name star_n}, _) $ _)) = + thm_of t + | thm_of_goal _ = raise MATCH; + in + SUBGOAL (fn (t, i) => + resolve_tac ctxt [thm_of_goal (strip_all_body t |> Logic.strip_imp_concl)] i + handle MATCH => no_tac) + end; +end; + fun transfer_thm_of ctxt ths t = let val {intros,unfolds,refolds,consts} = TransferData.get (Context.Proof ctxt); @@ -58,12 +77,12 @@ val unfolds' = map meta unfolds and refolds' = map meta refolds; val (_$_$t') = Thm.concl_of (Raw_Simplifier.rewrite ctxt true unfolds' (Thm.cterm_of ctxt t)) val u = unstar_term consts t' - val tac = + val tac = rewrite_goals_tac ctxt (ths' @ refolds' @ unfolds') THEN ALLGOALS (Object_Logic.full_atomize_tac ctxt) THEN match_tac ctxt [transitive_thm] 1 THEN resolve_tac ctxt [@{thm transfer_start}] 1 THEN - REPEAT_ALL_NEW (resolve_tac ctxt intros) 1 THEN + REPEAT_ALL_NEW (resolve_tac ctxt intros ORELSE' transfer_star_tac ctxt) 1 THEN match_tac ctxt [reflexive_thm] 1 in Goal.prove ctxt [] [] (Logic.mk_equals (t,u)) (K tac) end