2904
|
1 |
(* Title: HOL/Quot/PER0.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: Oscar Slotosch
|
|
4 |
Copyright 1997 Technische Universitaet Muenchen
|
|
5 |
|
|
6 |
*)
|
|
7 |
open PER0;
|
|
8 |
|
|
9 |
(* derive the characteristic axioms *)
|
|
10 |
val prems = goalw thy [per_def] "x === y ==> y === x";
|
|
11 |
by (cut_facts_tac prems 1);
|
3457
|
12 |
by (etac ax_per_sym 1);
|
2904
|
13 |
qed "per_sym";
|
|
14 |
|
|
15 |
val prems = goalw thy [per_def] "[| x === y; y === z |] ==> x === z";
|
|
16 |
by (cut_facts_tac prems 1);
|
3457
|
17 |
by (etac ax_per_trans 1);
|
|
18 |
by (assume_tac 1);
|
2904
|
19 |
qed "per_trans";
|
|
20 |
|
|
21 |
goalw thy [per_def] "(x::'a::er) === x";
|
3457
|
22 |
by (rtac ax_er_refl 1);
|
2904
|
23 |
qed "er_refl";
|
|
24 |
|
|
25 |
(* now everything works without axclasses *)
|
|
26 |
|
|
27 |
goal thy "x===y=y===x";
|
3457
|
28 |
by (rtac iffI 1);
|
|
29 |
by (etac per_sym 1);
|
|
30 |
by (etac per_sym 1);
|
2904
|
31 |
qed "per_sym2";
|
|
32 |
|
|
33 |
val prems = goal thy "x===y ==> x===x";
|
|
34 |
by (cut_facts_tac prems 1);
|
3457
|
35 |
by (rtac per_trans 1);by (assume_tac 1);
|
|
36 |
by (etac per_sym 1);
|
2904
|
37 |
qed "sym2refl1";
|
|
38 |
|
|
39 |
val prems = goal thy "x===y ==> y===y";
|
|
40 |
by (cut_facts_tac prems 1);
|
3457
|
41 |
by (rtac per_trans 1);by (assume_tac 2);
|
|
42 |
by (etac per_sym 1);
|
2904
|
43 |
qed "sym2refl2";
|
|
44 |
|
3058
|
45 |
val prems = goalw thy [Dom] "x:D ==> x === x";
|
2904
|
46 |
by (cut_facts_tac prems 1);
|
|
47 |
by (fast_tac set_cs 1);
|
|
48 |
qed "DomainD";
|
|
49 |
|
3058
|
50 |
val prems = goalw thy [Dom] "x === x ==> x:D";
|
2904
|
51 |
by (cut_facts_tac prems 1);
|
|
52 |
by (fast_tac set_cs 1);
|
|
53 |
qed "DomainI";
|
|
54 |
|
|
55 |
goal thy "x:D = x===x";
|
3457
|
56 |
by (rtac iffI 1);
|
|
57 |
by (etac DomainD 1);
|
|
58 |
by (etac DomainI 1);
|
2904
|
59 |
qed "DomainEq";
|
|
60 |
|
|
61 |
goal thy "(~ x === y) = (~ y === x)";
|
3457
|
62 |
by (rtac (per_sym2 RS arg_cong) 1);
|
2904
|
63 |
qed "per_not_sym";
|
|
64 |
|
|
65 |
(* show that PER work only on D *)
|
|
66 |
val prems = goal thy "x===y ==> x:D";
|
|
67 |
by (cut_facts_tac prems 1);
|
3457
|
68 |
by (etac (sym2refl1 RS DomainI) 1);
|
2904
|
69 |
qed "DomainI_left";
|
|
70 |
|
|
71 |
val prems = goal thy "x===y ==> y:D";
|
|
72 |
by (cut_facts_tac prems 1);
|
3457
|
73 |
by (etac (sym2refl2 RS DomainI) 1);
|
2904
|
74 |
qed "DomainI_right";
|
|
75 |
|
3058
|
76 |
val prems = goalw thy [Dom] "x~:D ==> ~x===y";
|
2904
|
77 |
by (cut_facts_tac prems 1);
|
3457
|
78 |
by (res_inst_tac [("Q","x===y")] (excluded_middle RS disjE) 1);by (assume_tac 1);
|
|
79 |
by (dtac sym2refl1 1);
|
2904
|
80 |
by (fast_tac set_cs 1);
|
|
81 |
qed "notDomainE1";
|
|
82 |
|
3058
|
83 |
val prems = goalw thy [Dom] "x~:D ==> ~y===x";
|
2904
|
84 |
by (cut_facts_tac prems 1);
|
3457
|
85 |
by (res_inst_tac [("Q","y===x")] (excluded_middle RS disjE) 1);by (assume_tac 1);
|
|
86 |
by (dtac sym2refl2 1);
|
2904
|
87 |
by (fast_tac set_cs 1);
|
|
88 |
qed "notDomainE2";
|
|
89 |
|
|
90 |
(* theorems for equivalence relations *)
|
|
91 |
goal thy "(x::'a::er) : D";
|
3457
|
92 |
by (rtac DomainI 1);
|
|
93 |
by (rtac er_refl 1);
|
2904
|
94 |
qed "er_Domain";
|
|
95 |
|
|
96 |
(* witnesses for "=>" ::(per,per)per *)
|
|
97 |
val prems = goalw thy [fun_per_def]"eqv (x::'a::per => 'b::per) y ==> eqv y x";
|
|
98 |
by (cut_facts_tac prems 1);
|
|
99 |
by (safe_tac HOL_cs);
|
|
100 |
by (asm_full_simp_tac (HOL_ss addsimps [per_sym2]) 1);
|
|
101 |
qed "per_sym_fun";
|
|
102 |
|
|
103 |
val prems = goalw thy [fun_per_def]
|
|
104 |
"[| eqv (f::'a::per=>'b::per) g;eqv g h|] ==> eqv f h";
|
|
105 |
by (cut_facts_tac prems 1);
|
|
106 |
by (safe_tac HOL_cs);
|
|
107 |
by (REPEAT (dtac spec 1));
|
|
108 |
by (res_inst_tac [("y","g y")] per_trans 1);
|
3457
|
109 |
by (rtac mp 1);by (assume_tac 1);
|
2904
|
110 |
by (Asm_simp_tac 1);
|
3457
|
111 |
by (rtac mp 1);by (assume_tac 1);
|
4089
|
112 |
by (asm_simp_tac (simpset() addsimps [sym2refl2]) 1);
|
2904
|
113 |
qed "per_trans_fun";
|
|
114 |
|
|
115 |
|