| 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 *)
 | 
| 5858 |     10 | Goalw [per_def] "x === y ==> y === x";
 | 
| 3457 |     11 | by (etac ax_per_sym 1);
 | 
| 2904 |     12 | qed "per_sym";
 | 
|  |     13 | 
 | 
| 5858 |     14 | Goalw [per_def] "[| x === y; y === z |] ==> x === z";
 | 
| 3457 |     15 | by (etac ax_per_trans 1);
 | 
|  |     16 | by (assume_tac 1);
 | 
| 2904 |     17 | qed "per_trans";
 | 
|  |     18 | 
 | 
| 5069 |     19 | Goalw [per_def] "(x::'a::er) === x";
 | 
| 3457 |     20 | by (rtac ax_er_refl 1);
 | 
| 2904 |     21 | qed "er_refl";
 | 
|  |     22 | 
 | 
|  |     23 | (* now everything works without axclasses *)
 | 
|  |     24 | 
 | 
| 5069 |     25 | Goal "x===y=y===x";
 | 
| 3457 |     26 | by (rtac iffI 1);
 | 
|  |     27 | by (etac per_sym 1);
 | 
|  |     28 | by (etac per_sym 1);
 | 
| 2904 |     29 | qed "per_sym2";
 | 
|  |     30 | 
 | 
| 5858 |     31 | Goal "x===y ==> x===x";
 | 
| 3457 |     32 | by (rtac per_trans 1);by (assume_tac 1);
 | 
|  |     33 | by (etac per_sym 1);
 | 
| 2904 |     34 | qed "sym2refl1";
 | 
|  |     35 | 
 | 
| 5858 |     36 | Goal "x===y ==> y===y";
 | 
| 3457 |     37 | by (rtac per_trans 1);by (assume_tac 2);
 | 
|  |     38 | by (etac per_sym 1);
 | 
| 2904 |     39 | qed "sym2refl2";
 | 
|  |     40 | 
 | 
| 5858 |     41 | Goalw [Dom] "x:D ==> x === x";
 | 
|  |     42 | by (Blast_tac 1);
 | 
| 2904 |     43 | qed "DomainD";
 | 
|  |     44 | 
 | 
| 5858 |     45 | Goalw [Dom] "x === x ==> x:D";
 | 
|  |     46 | by (Blast_tac 1);
 | 
| 2904 |     47 | qed "DomainI";
 | 
|  |     48 | 
 | 
| 5069 |     49 | Goal "x:D = x===x";
 | 
| 3457 |     50 | by (rtac iffI 1);
 | 
|  |     51 | by (etac DomainD 1);
 | 
|  |     52 | by (etac DomainI 1);
 | 
| 2904 |     53 | qed "DomainEq";
 | 
|  |     54 | 
 | 
| 5069 |     55 | Goal "(~ x === y) = (~ y === x)";
 | 
| 3457 |     56 | by (rtac (per_sym2 RS arg_cong) 1);
 | 
| 2904 |     57 | qed "per_not_sym";
 | 
|  |     58 | 
 | 
|  |     59 | (* show that PER work only on D *)
 | 
| 5858 |     60 | Goal "x===y ==> x:D";
 | 
| 3457 |     61 | by (etac (sym2refl1 RS DomainI) 1);
 | 
| 2904 |     62 | qed "DomainI_left"; 
 | 
|  |     63 | 
 | 
| 5858 |     64 | Goal "x===y ==> y:D";
 | 
| 3457 |     65 | by (etac (sym2refl2 RS DomainI) 1);
 | 
| 2904 |     66 | qed "DomainI_right"; 
 | 
|  |     67 | 
 | 
| 5858 |     68 | Goalw [Dom] "x~:D ==> ~x===y";
 | 
|  |     69 | by (res_inst_tac [("Q","x===y")] (excluded_middle RS disjE) 1);
 | 
|  |     70 | by (assume_tac 1);
 | 
| 3457 |     71 | by (dtac sym2refl1 1);
 | 
| 5858 |     72 | by (Blast_tac 1);
 | 
| 2904 |     73 | qed "notDomainE1"; 
 | 
|  |     74 | 
 | 
| 5858 |     75 | Goalw [Dom] "x~:D ==> ~y===x";
 | 
|  |     76 | by (res_inst_tac [("Q","y===x")] (excluded_middle RS disjE) 1);
 | 
|  |     77 | by (assume_tac 1);
 | 
| 3457 |     78 | by (dtac sym2refl2 1);
 | 
| 5858 |     79 | by (Blast_tac 1);
 | 
| 2904 |     80 | qed "notDomainE2"; 
 | 
|  |     81 | 
 | 
|  |     82 | (* theorems for equivalence relations *)
 | 
| 5069 |     83 | Goal "(x::'a::er) : D";
 | 
| 3457 |     84 | by (rtac DomainI 1);
 | 
|  |     85 | by (rtac er_refl 1);
 | 
| 2904 |     86 | qed "er_Domain";
 | 
|  |     87 | 
 | 
|  |     88 | (* witnesses for "=>" ::(per,per)per  *)
 | 
| 5858 |     89 | Goalw [fun_per_def]"eqv (x::'a::per => 'b::per) y ==> eqv y x";
 | 
|  |     90 | by (auto_tac (claset(), simpset() addsimps [per_sym2]));
 | 
| 2904 |     91 | qed "per_sym_fun";
 | 
|  |     92 | 
 | 
| 5858 |     93 | Goalw [fun_per_def] "[| eqv (f::'a::per=>'b::per) g;eqv g h|] ==> eqv f h";
 | 
|  |     94 | by Safe_tac;
 | 
| 2904 |     95 | by (REPEAT (dtac spec 1));
 | 
|  |     96 | by (res_inst_tac [("y","g y")] per_trans 1);
 | 
| 3457 |     97 | by (rtac mp 1);by (assume_tac 1);
 | 
| 2904 |     98 | by (Asm_simp_tac 1);
 | 
| 3457 |     99 | by (rtac mp 1);by (assume_tac 1);
 | 
| 4089 |    100 | by (asm_simp_tac (simpset() addsimps [sym2refl2]) 1);
 | 
| 2904 |    101 | qed "per_trans_fun";
 | 
|  |    102 | 
 | 
|  |    103 | 
 |