doc-src/TutorialI/Sets/Relations.thy
author noschinl
Wed, 08 Dec 2010 18:18:36 +0100
changeset 41826 18d4d2b60016
parent 36745 403585a89772
child 42637 381fdcab0f36
permissions -rw-r--r--
introduce attribute case_prod for combining case rules * * * [PATCH 1/9] Make tac independent of consumes From 1a53ef4c070f924ff8e69529b0c1b13fa2844722 Mon Sep 17 00:00:00 2001 --- case_product.ML | 11 ++++++----- 1 files changed, 6 insertions(+), 5 deletions(-) * * * [PATCH 2/9] Replace MRS by OF From da55d08ae287bfdc18dec554067b45a3e298c243 Mon Sep 17 00:00:00 2001 --- case_product.ML | 4 ++-- 1 files changed, 2 insertions(+), 2 deletions(-) * * * [PATCH 3/9] Clean up tactic From d26d50caaa3526c8c0625d7395467c6914f1a8d9 Mon Sep 17 00:00:00 2001 --- case_product.ML | 13 +++++-------- 1 files changed, 5 insertions(+), 8 deletions(-) * * * [PATCH 4/9] Move out get_consumes a bit more From 6ed5415f29cc43cea30c216edb1e74ec6c0f6c33 Mon Sep 17 00:00:00 2001 --- case_product.ML | 4 +++- 1 files changed, 3 insertions(+), 1 deletions(-) * * * [PATCH 5/9] Clean up tactic From d301cfe6377e9f7db74b190fb085e0e66eeadfb5 Mon Sep 17 00:00:00 2001 --- case_product.ML | 3 +-- 1 files changed, 1 insertions(+), 2 deletions(-) * * * [PATCH 6/9] Cleanup build_concl_prems From c30889e131e74a92caac5b1e6d3d816890406ca7 Mon Sep 17 00:00:00 2001 --- case_product.ML | 18 ++++++++---------- 1 files changed, 8 insertions(+), 10 deletions(-) * * * [PATCH 7/9] Split logic & annotation a bit From e065606118308c96e013fad72ab9ad89b5a4b945 Mon Sep 17 00:00:00 2001 --- case_product.ML | 26 ++++++++++++++++++-------- 1 files changed, 18 insertions(+), 8 deletions(-) * * * [PATCH 8/9] Remove dependency on consumes attribute From e2a4de044481586d6127bbabd0ef48d0e3ab57c0 Mon Sep 17 00:00:00 2001 --- case_product.ML | 46 ++++++++++++++++++++++------------------------ 1 files changed, 22 insertions(+), 24 deletions(-) * * * [PATCH 9/9] whitespace From 59f5036bd8f825da4a362970292a34ec06c0f1a2 Mon Sep 17 00:00:00 2001 --- case_product.ML | 2 +- 1 files changed, 1 insertions(+), 1 deletions(-)

(* ID:         $Id$ *)
theory Relations imports Main begin

ML "Pretty.margin_default := 64"

(*Id is only used in UNITY*)
(*refl, antisym,trans,univalent,\<dots> ho hum*)

text{*
@{thm[display] Id_def[no_vars]}
\rulename{Id_def}
*}

text{*
@{thm[display] rel_comp_def[no_vars]}
\rulename{rel_comp_def}
*}

text{*
@{thm[display] R_O_Id[no_vars]}
\rulename{R_O_Id}
*}

text{*
@{thm[display] rel_comp_mono[no_vars]}
\rulename{rel_comp_mono}
*}

text{*
@{thm[display] converse_iff[no_vars]}
\rulename{converse_iff}
*}

text{*
@{thm[display] converse_rel_comp[no_vars]}
\rulename{converse_rel_comp}
*}

text{*
@{thm[display] Image_iff[no_vars]}
\rulename{Image_iff}
*}

text{*
@{thm[display] Image_UN[no_vars]}
\rulename{Image_UN}
*}

text{*
@{thm[display] Domain_iff[no_vars]}
\rulename{Domain_iff}
*}

text{*
@{thm[display] Range_iff[no_vars]}
\rulename{Range_iff}
*}

text{*
@{thm[display] relpow.simps[no_vars]}
\rulename{relpow.simps}

@{thm[display] rtrancl_refl[no_vars]}
\rulename{rtrancl_refl}

@{thm[display] r_into_rtrancl[no_vars]}
\rulename{r_into_rtrancl}

@{thm[display] rtrancl_trans[no_vars]}
\rulename{rtrancl_trans}

@{thm[display] rtrancl_induct[no_vars]}
\rulename{rtrancl_induct}

@{thm[display] rtrancl_idemp[no_vars]}
\rulename{rtrancl_idemp}

@{thm[display] r_into_trancl[no_vars]}
\rulename{r_into_trancl}

@{thm[display] trancl_trans[no_vars]}
\rulename{trancl_trans}

@{thm[display] trancl_into_rtrancl[no_vars]}
\rulename{trancl_into_rtrancl}

@{thm[display] trancl_converse[no_vars]}
\rulename{trancl_converse}
*}

text{*Relations.  transitive closure*}

lemma rtrancl_converseD: "(x,y) \<in> (r\<inverse>)\<^sup>* \<Longrightarrow> (y,x) \<in> r\<^sup>*"
apply (erule rtrancl_induct)
txt{*
@{subgoals[display,indent=0,margin=65]}
*};
 apply (rule rtrancl_refl)
apply (blast intro: rtrancl_trans)
done


lemma rtrancl_converseI: "(y,x) \<in> r\<^sup>* \<Longrightarrow> (x,y) \<in> (r\<inverse>)\<^sup>*"
apply (erule rtrancl_induct)
 apply (rule rtrancl_refl)
apply (blast intro: rtrancl_trans)
done

lemma rtrancl_converse: "(r\<inverse>)\<^sup>* = (r\<^sup>*)\<inverse>"
by (auto intro: rtrancl_converseI dest: rtrancl_converseD)

lemma rtrancl_converse: "(r\<inverse>)\<^sup>* = (r\<^sup>*)\<inverse>"
apply (intro equalityI subsetI)
txt{*
after intro rules

@{subgoals[display,indent=0,margin=65]}
*};
apply clarify
txt{*
after splitting
@{subgoals[display,indent=0,margin=65]}
*};
oops


lemma "(\<forall>u v. (u,v) \<in> A \<longrightarrow> u=v) \<Longrightarrow> A \<subseteq> Id"
apply (rule subsetI)
txt{*
@{subgoals[display,indent=0,margin=65]}

after subsetI
*};
apply clarify
txt{*
@{subgoals[display,indent=0,margin=65]}

subgoals after clarify
*};
by blast




text{*rejects*}

lemma "(a \<in> {z. P z} \<union> {y. Q y}) = P a \<or> Q a"
apply (blast)
done

text{*Pow, Inter too little used*}

lemma "(A \<subset> B) = (A \<subseteq> B \<and> A \<noteq> B)"
apply (simp add: psubset_eq)
done

end