author | huffman |
Wed, 18 Apr 2012 15:09:07 +0200 | |
changeset 47538 | 1f0ec5b8135a |
parent 47537 | b06be48923a4 |
child 47539 | 436ae5ea4f80 |
child 47544 | e455cdaac479 |
--- a/src/HOL/Lifting.thy Wed Apr 18 14:59:04 2012 +0200 +++ b/src/HOL/Lifting.thy Wed Apr 18 15:09:07 2012 +0200 @@ -296,6 +296,10 @@ lemma Quotient_rel_eq_transfer: "(T ===> T ===> op =) R (op =)" using 1 unfolding Quotient_alt_def fun_rel_def by simp +lemma Quotient_abs_induct: + assumes "\<And>y. R y y \<Longrightarrow> P (Abs y)" shows "P x" + using 1 assms unfolding Quotient_def by metis + end text {* Generating transfer rules for total quotients. *}