# HG changeset patch # User paulson # Date 1614610011 0 # Node ID 00e0f7724c06b76189ae01c744581ff9218536fe # Parent 2aef2de6b17cb8f71f8e9bebfdb7bb6854b7a658 tiny bit of lemma hacking diff -r 2aef2de6b17c -r 00e0f7724c06 src/HOL/Library/FuncSet.thy --- a/src/HOL/Library/FuncSet.thy Mon Mar 01 08:16:22 2021 +0100 +++ b/src/HOL/Library/FuncSet.thy Mon Mar 01 14:46:51 2021 +0000 @@ -175,9 +175,6 @@ lemma restrict_cong: "I = J \ (\i. i \ J =simp=> f i = g i) \ restrict f I = restrict g J" by (auto simp: restrict_def fun_eq_iff simp_implies_def) -lemma restrict_in_funcset: "(\x. x \ A \ f x \ B) \ (\x\A. f x) \ A \ B" - by (simp add: Pi_def restrict_def) - lemma restrictI[intro!]: "(\x. x \ A \ f x \ B x) \ (\x\A. f x) \ Pi A B" by (simp add: Pi_def restrict_def) @@ -435,6 +432,9 @@ lemma PiE_iff: "f \ Pi\<^sub>E I X \ (\i\I. f i \ X i) \ f \ extensional I" by (simp add: PiE_def Pi_iff) +lemma restrict_PiE_iff [simp]: "restrict f I \ Pi\<^sub>E I X \ (\i \ I. f i \ X i)" + by (simp add: PiE_iff) + lemma ext_funcset_to_sing_iff [simp]: "A \\<^sub>E {a} = {\x\A. a}" by (auto simp: PiE_def Pi_iff extensionalityI) @@ -495,10 +495,7 @@ lemma extensional_funcset_fun_upd_restricts_rangeI: "\y \ S. f x \ f y \ f \ (insert x S) \\<^sub>E T \ f(x := undefined) \ S \\<^sub>E (T - {f x})" unfolding extensional_funcset_def extensional_def - apply auto - apply (case_tac "x = xa") - apply auto - done + by (auto split: if_split_asm) lemma extensional_funcset_fun_upd_extends_rangeI: assumes "a \ T" "f \ S \\<^sub>E (T - {a})" diff -r 2aef2de6b17c -r 00e0f7724c06 src/HOL/Metis_Examples/Tarski.thy --- a/src/HOL/Metis_Examples/Tarski.thy Mon Mar 01 08:16:22 2021 +0100 +++ b/src/HOL/Metis_Examples/Tarski.thy Mon Mar 01 14:46:51 2021 +0000 @@ -905,8 +905,8 @@ lemma (in Tarski) intY1_func: "(\x \ intY1. f x) \ intY1 \ intY1" -apply (rule restrict_in_funcset) -apply (metis intY1_f_closed restrict_in_funcset) +apply (rule restrictI) +apply (metis intY1_f_closed) done @@ -1018,15 +1018,8 @@ theorem (in CLF) Tarski_full: "(| pset = P, order = induced P r|) \ CompleteLattice" -(*sledgehammer*) -apply (rule CompleteLatticeI_simp) -apply (rule fixf_po, clarify) -(*never proved, 2007-01-22*) + using A_def CLF_axioms P_def Tarski.intro Tarski_axioms.intro fixf_po r_def by blast (*sledgehammer*) -apply (simp add: P_def A_def r_def) -apply (blast intro!: Tarski.tarski_full_lemma [OF Tarski.intro, OF CLF.intro Tarski_axioms.intro, - OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro] cl_po cl_co f_cl) -done declare (in CLF) fixf_po [rule del] P_def [simp del] A_def [simp del] r_def [simp del] Tarski.tarski_full_lemma [rule del] cl_po [rule del] cl_co [rule del]