diff -r eee04bbbae7e -r 6f8f9835e219 src/HOL/Predicate.thy --- a/src/HOL/Predicate.thy Fri Dec 11 14:43:55 2009 +0100 +++ b/src/HOL/Predicate.thy Fri Dec 11 14:43:56 2009 +0100 @@ -19,6 +19,53 @@ subsection {* Predicates as (complete) lattices *} + +text {* + Handy introduction and elimination rules for @{text "\"} + on unary and binary predicates +*} + +lemma predicate1I: + assumes PQ: "\x. P x \ Q x" + shows "P \ Q" + apply (rule le_funI) + apply (rule le_boolI) + apply (rule PQ) + apply assumption + done + +lemma predicate1D [Pure.dest?, dest?]: + "P \ Q \ P x \ Q x" + apply (erule le_funE) + apply (erule le_boolE) + apply assumption+ + done + +lemma rev_predicate1D: + "P x ==> P <= Q ==> Q x" + by (rule predicate1D) + +lemma predicate2I [Pure.intro!, intro!]: + assumes PQ: "\x y. P x y \ Q x y" + shows "P \ Q" + apply (rule le_funI)+ + apply (rule le_boolI) + apply (rule PQ) + apply assumption + done + +lemma predicate2D [Pure.dest, dest]: + "P \ Q \ P x y \ Q x y" + apply (erule le_funE)+ + apply (erule le_boolE) + apply assumption+ + done + +lemma rev_predicate2D: + "P x y ==> P <= Q ==> Q x y" + by (rule predicate2D) + + subsubsection {* Equality *} lemma pred_equals_eq: "((\x. x \ R) = (\x. x \ S)) = (R = S)"