author | Christian Urban <urbanc@in.tum.de> |
Mon, 12 Oct 2009 23:06:14 +0200 | |
changeset 79 | c0c41fefeb06 |
parent 57 | 13be92f5b638 |
child 83 | e8f352546ad8 |
permissions | -rw-r--r-- |
6 | 1 |
theory Prove |
79
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
2 |
imports Main QuotScript QuotList |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
3 |
uses ("quotient.ML") |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
4 |
begin |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
5 |
|
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
6 |
locale QUOT_TYPE = |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
7 |
fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
8 |
and Abs :: "('a \<Rightarrow> bool) \<Rightarrow> 'b" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
9 |
and Rep :: "'b \<Rightarrow> ('a \<Rightarrow> bool)" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
10 |
assumes equiv: "EQUIV R" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
11 |
and rep_prop: "\<And>y. \<exists>x. Rep y = R x" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
12 |
and rep_inverse: "\<And>x. Abs (Rep x) = x" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
13 |
and abs_inverse: "\<And>x. (Rep (Abs (R x))) = (R x)" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
14 |
and rep_inject: "\<And>x y. (Rep x = Rep y) = (x = y)" |
6 | 15 |
begin |
16 |
||
79
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
17 |
definition |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
18 |
"ABS x \<equiv> Abs (R x)" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
19 |
|
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
20 |
definition |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
21 |
"REP a = Eps (Rep a)" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
22 |
|
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
23 |
lemma lem9: |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
24 |
shows "R (Eps (R x)) = R x" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
25 |
proof - |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
26 |
have a: "R x x" using equiv by (simp add: EQUIV_REFL_SYM_TRANS REFL_def) |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
27 |
then have "R x (Eps (R x))" by (rule someI) |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
28 |
then show "R (Eps (R x)) = R x" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
29 |
using equiv unfolding EQUIV_def by simp |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
30 |
qed |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
31 |
|
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
32 |
theorem thm10: |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
33 |
shows "ABS (REP a) \<equiv> a" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
34 |
apply (rule eq_reflection) |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
35 |
unfolding ABS_def REP_def |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
36 |
proof - |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
37 |
from rep_prop |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
38 |
obtain x where eq: "Rep a = R x" by auto |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
39 |
have "Abs (R (Eps (Rep a))) = Abs (R (Eps (R x)))" using eq by simp |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
40 |
also have "\<dots> = Abs (R x)" using lem9 by simp |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
41 |
also have "\<dots> = Abs (Rep a)" using eq by simp |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
42 |
also have "\<dots> = a" using rep_inverse by simp |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
43 |
finally |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
44 |
show "Abs (R (Eps (Rep a))) = a" by simp |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
45 |
qed |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
46 |
|
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
47 |
lemma REP_refl: |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
48 |
shows "R (REP a) (REP a)" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
49 |
unfolding REP_def |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
50 |
by (simp add: equiv[simplified EQUIV_def]) |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
51 |
|
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
52 |
lemma lem7: |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
53 |
shows "(R x = R y) = (Abs (R x) = Abs (R y))" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
54 |
apply(rule iffI) |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
55 |
apply(simp) |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
56 |
apply(drule rep_inject[THEN iffD2]) |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
57 |
apply(simp add: abs_inverse) |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
58 |
done |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
59 |
|
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
60 |
theorem thm11: |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
61 |
shows "R r r' = (ABS r = ABS r')" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
62 |
unfolding ABS_def |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
63 |
by (simp only: equiv[simplified EQUIV_def] lem7) |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
64 |
|
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
65 |
|
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
66 |
lemma REP_ABS_rsp: |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
67 |
shows "R f (REP (ABS g)) = R f g" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
68 |
and "R (REP (ABS g)) f = R g f" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
69 |
by (simp_all add: thm10 thm11) |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
70 |
|
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
71 |
lemma QUOTIENT: |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
72 |
"QUOTIENT R ABS REP" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
73 |
apply(unfold QUOTIENT_def) |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
74 |
apply(simp add: thm10) |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
75 |
apply(simp add: REP_refl) |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
76 |
apply(subst thm11[symmetric]) |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
77 |
apply(simp add: equiv[simplified EQUIV_def]) |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
78 |
done |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
79 |
|
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
80 |
lemma R_trans: |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
81 |
assumes ab: "R a b" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
82 |
and bc: "R b c" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
83 |
shows "R a c" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
84 |
proof - |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
85 |
have tr: "TRANS R" using equiv EQUIV_REFL_SYM_TRANS[of R] by simp |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
86 |
moreover have ab: "R a b" by fact |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
87 |
moreover have bc: "R b c" by fact |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
88 |
ultimately show "R a c" unfolding TRANS_def by blast |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
89 |
qed |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
90 |
|
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
91 |
lemma R_sym: |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
92 |
assumes ab: "R a b" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
93 |
shows "R b a" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
94 |
proof - |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
95 |
have re: "SYM R" using equiv EQUIV_REFL_SYM_TRANS[of R] by simp |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
96 |
then show "R b a" using ab unfolding SYM_def by blast |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
97 |
qed |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
98 |
|
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
99 |
lemma R_trans2: |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
100 |
assumes ac: "R a c" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
101 |
and bd: "R b d" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
102 |
shows "R a b = R c d" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
103 |
proof |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
104 |
assume "R a b" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
105 |
then have "R b a" using R_sym by blast |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
106 |
then have "R b c" using ac R_trans by blast |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
107 |
then have "R c b" using R_sym by blast |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
108 |
then show "R c d" using bd R_trans by blast |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
109 |
next |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
110 |
assume "R c d" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
111 |
then have "R a d" using ac R_trans by blast |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
112 |
then have "R d a" using R_sym by blast |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
113 |
then have "R b a" using bd R_trans by blast |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
114 |
then show "R a b" using R_sym by blast |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
115 |
qed |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
116 |
|
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
117 |
lemma REPS_same: |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
118 |
shows "R (REP a) (REP b) \<equiv> (a = b)" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
119 |
proof - |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
120 |
have "R (REP a) (REP b) = (a = b)" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
121 |
proof |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
122 |
assume as: "R (REP a) (REP b)" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
123 |
from rep_prop |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
124 |
obtain x y |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
125 |
where eqs: "Rep a = R x" "Rep b = R y" by blast |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
126 |
from eqs have "R (Eps (R x)) (Eps (R y))" using as unfolding REP_def by simp |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
127 |
then have "R x (Eps (R y))" using lem9 by simp |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
128 |
then have "R (Eps (R y)) x" using R_sym by blast |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
129 |
then have "R y x" using lem9 by simp |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
130 |
then have "R x y" using R_sym by blast |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
131 |
then have "ABS x = ABS y" using thm11 by simp |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
132 |
then have "Abs (Rep a) = Abs (Rep b)" using eqs unfolding ABS_def by simp |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
133 |
then show "a = b" using rep_inverse by simp |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
134 |
next |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
135 |
assume ab: "a = b" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
136 |
have "REFL R" using equiv EQUIV_REFL_SYM_TRANS[of R] by simp |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
137 |
then show "R (REP a) (REP b)" unfolding REFL_def using ab by auto |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
138 |
qed |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
139 |
then show "R (REP a) (REP b) \<equiv> (a = b)" by simp |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
140 |
qed |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
141 |
|
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
142 |
end |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
143 |
|
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
144 |
use "quotient.ML" |
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
Christian Urban <urbanc@in.tum.de>
parents:
57
diff
changeset
|
145 |
|
6 | 146 |
ML {* |
147 |
val r = ref (NONE:(unit -> term) option) |
|
148 |
*} |
|
46 | 149 |
|
6 | 150 |
ML {* |
151 |
let |
|
48 | 152 |
fun after_qed thm_name thms lthy = |
153 |
LocalTheory.note Thm.theoremK (thm_name, (flat thms)) lthy |> snd |
|
154 |
||
155 |
fun setup_proof (name_spec, (txt, pos)) lthy = |
|
156 |
let |
|
157 |
val trm = ML_Context.evaluate lthy true ("r", r) txt |
|
158 |
in |
|
159 |
Proof.theorem_i NONE (after_qed name_spec) [[(trm,[])]] lthy |
|
57 | 160 |
end |
6 | 161 |
|
48 | 162 |
val parser = SpecParse.opt_thm_name ":" -- OuterParse.ML_source |
6 | 163 |
in |
46 | 164 |
OuterSyntax.local_theory_to_proof "prove" "proving a proposition" |
48 | 165 |
OuterKeyword.thy_goal (parser >> setup_proof) |
6 | 166 |
end |
167 |
*} |
|
168 |
||
169 |
end |