diff -r 05e5d68c9627 -r f1be8028a4a9 Nominal/activities/tphols09/IDW/ExampleConstruction.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Nominal/activities/tphols09/IDW/ExampleConstruction.thy Wed Mar 30 17:27:34 2016 +0100 @@ -0,0 +1,153 @@ +theory IndExamples +imports Main +begin + +section {* Transitive Closure *} + +text {* + Introduction rules: + @{term "trcl R x x"} + @{term "R x y \ trcl R y z \ trcl R x z"} +*} + +definition "trcl \ \R x y. + \P. (\x. P x x) \ (\x y z. R x y \ P y z \ P x z) \ P x y" + +lemma trcl_induct: + assumes trcl: "trcl R x y" + shows "(\x. P x x) \ (\x y z. R x y \ P y z \ P x z) \ P x y" + apply (atomize (full)) + apply (cut_tac trcl) + apply (unfold trcl_def) + apply (drule spec [where x=P]) + apply assumption + done + +lemma trcl_base: "trcl R x x" + apply (unfold trcl_def) + apply (rule allI impI)+ + apply (drule spec) + apply assumption + done + +lemma trcl_step: "R x y \ trcl R y z \ trcl R x z" + apply (unfold trcl_def) + apply (rule allI impI)+ + proof - + case goal1 + show ?case + apply (rule goal1(4) [rule_format]) + apply (rule goal1(1)) + apply (rule goal1(2) [THEN spec [where x=P], THEN mp, THEN mp, + OF goal1(3-4)]) + done + qed + + + +section {* Even and Odd Numbers, Mutually Inductive *} + +text {* + Introduction rules: + @{term "even 0"} + @{term "odd m \ even (Suc m)"} + @{term "even m \ odd (Suc m)"} +*} + +definition "even \ \n. + \P Q. P 0 \ (\m. Q m \ P (Suc m)) \ (\m. P m \ Q (Suc m)) \ P n" + +definition "odd \ \n. + \P Q. P 0 \ (\m. Q m \ P (Suc m)) \ (\m. P m \ Q (Suc m)) \ Q n" + +lemma even_induct: + assumes even: "even n" + shows "P 0 \ (\m. Q m \ P (Suc m)) \ (\m. P m \ Q (Suc m)) \ P n" + apply (atomize (full)) + apply (cut_tac even) + apply (unfold even_def) + apply (drule spec [where x=P]) + apply (drule spec [where x=Q]) + apply assumption + done + +lemma odd_induct: + assumes odd: "odd n" + shows "P 0 \ (\m. Q m \ P (Suc m)) \ (\m. P m \ Q (Suc m)) \ Q n" + apply (atomize (full)) + apply (cut_tac odd) + apply (unfold odd_def) + apply (drule spec [where x=P]) + apply (drule spec [where x=Q]) + apply assumption + done + +lemma even0: "even 0" + apply (unfold even_def) + apply (rule allI impI)+ + apply assumption + done + +lemma evenS: "odd m \ even (Suc m)" + apply (unfold odd_def even_def) + apply (rule allI impI)+ + proof - + case goal1 + show ?case + apply (rule goal1(3) [rule_format]) + apply (rule goal1(1) [THEN spec [where x=P], THEN spec [where x=Q], + THEN mp, THEN mp, THEN mp, OF goal1(2-4)]) + done + qed + +lemma oddS: "even m \ odd (Suc m)" + apply (unfold odd_def even_def) + apply (rule allI impI)+ + proof - + case goal1 + show ?case + apply (rule goal1(4) [rule_format]) + apply (rule goal1(1) [THEN spec [where x=P], THEN spec [where x=Q], + THEN mp, THEN mp, THEN mp, OF goal1(2-4)]) + done + qed + + + +section {* Accessible Part *} + +text {* + Introduction rules: + @{term "(\y. R y x \ accpart R y) \ accpart R x"} +*} + +definition "accpart \ \R x. \P. (\x. (\y. R y x \ P y) \ P x) \ P x" + +lemma accpart_induct: + assumes acc: "accpart R x" + shows "(\x. (\y. R y x \ P y) \ P x) \ P x" + apply (atomize (full)) + apply (cut_tac acc) + apply (unfold accpart_def) + apply (drule spec [where x=P]) + apply assumption + done + +lemma accpartI: "(\y. R y x \ accpart R y) \ accpart R x" + apply (unfold accpart_def) + apply (rule allI impI)+ + proof - + case goal1 + note goal1' = this + show ?case + apply (rule goal1'(2) [rule_format]) + proof - + case goal1 + show ?case + apply (rule goal1'(1) [OF goal1, THEN spec [where x=P], + THEN mp, OF goal1'(2)]) + done + qed + qed + +end