| author | wenzelm | 
| Fri, 19 Dec 1997 10:18:58 +0100 | |
| changeset 4447 | b7ee449eb345 | 
| parent 3842 | b55686a7b22c | 
| child 4477 | b3e5857d8d99 | 
| permissions | -rw-r--r-- | 
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 1 | (* Title: HOL/Quot/HQUOT.ML | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 2 | ID: $Id$ | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 3 | Author: Oscar Slotosch | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 4 | Copyright 1997 Technische Universitaet Muenchen | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 5 | |
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 6 | *) | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 7 | open HQUOT; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 8 | |
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 9 | (* first prove some helpful lemmas *) | 
| 3842 | 10 | goalw thy [quot_def] "{y. y===x}:quot";
 | 
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 11 | by (Asm_simp_tac 1); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 12 | by (fast_tac (set_cs addIs [per_sym]) 1); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 13 | qed "per_class_rep_quot"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 14 | |
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 15 | val prems = goal thy "Rep_quot a = Rep_quot b ==> a=b"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 16 | by (cut_facts_tac prems 1); | 
| 3457 | 17 | by (rtac (Rep_quot_inverse RS subst) 1); | 
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 18 | by (res_inst_tac [("t","a")] (Rep_quot_inverse RS subst) 1);
 | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 19 | by (Asm_simp_tac 1); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 20 | qed "quot_eq"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 21 | |
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 22 | (* prepare induction and exhaustiveness *) | 
| 3842 | 23 | val prems = goal thy "!s. s:quot --> P (Abs_quot s) ==> P x"; | 
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 24 | by (cut_facts_tac prems 1); | 
| 3457 | 25 | by (rtac (Abs_quot_inverse RS subst) 1); | 
| 26 | by (rtac Rep_quot 1); | |
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 27 | by (dres_inst_tac [("x","Rep_quot x")] spec 1);
 | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 28 | by (asm_full_simp_tac (HOL_ss addsimps [Rep_quot,Rep_quot_inverse]) 1); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 29 | qed "all_q"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 30 | |
| 3842 | 31 | goal thy "? s. s:quot & x=Abs_quot s"; | 
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 32 | by (res_inst_tac [("x","Rep_quot x")] exI 1);
 | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 33 | by (asm_full_simp_tac (HOL_ss addsimps [Rep_quot,Rep_quot_inverse]) 1); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 34 | qed "exh_q"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 35 | |
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 36 | (* some lemmas for the equivalence classes *) | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 37 | |
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 38 | (* equality and symmetry for equivalence classes *) | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 39 | val prems = goalw thy [peclass_def] "x===y==><[x]>=<[y]>"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 40 | by (cut_facts_tac prems 1); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 41 | by (res_inst_tac [("f","Abs_quot")] arg_cong 1);
 | 
| 3457 | 42 | by (rtac Collect_cong 1); | 
| 43 | by (rtac iffI 1); | |
| 44 | by (etac per_trans 1);by (assume_tac 1); | |
| 45 | by (etac per_trans 1); | |
| 46 | by (etac per_sym 1); | |
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 47 | qed "per_class_eqI"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 48 | |
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 49 | val prems = goalw thy [peclass_def] "<[(x::'a::er)]>=<[y]>==>x===y"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 50 | by (cut_facts_tac prems 1); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 51 | by (dres_inst_tac [("f","Rep_quot")] arg_cong 1);
 | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 52 | by (asm_full_simp_tac(HOL_ss addsimps [Abs_quot_inverse,per_class_rep_quot])1); | 
| 3457 | 53 | by (dres_inst_tac [("c","x")] equalityCE 1);by (assume_tac 3);
 | 
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 54 | by (safe_tac set_cs); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 55 | by (fast_tac (set_cs addIs [er_refl]) 1); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 56 | qed "er_class_eqE"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 57 | |
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 58 | val prems = goalw thy [peclass_def] "[|x:D;<[x]>=<[y]>|]==>x===y"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 59 | by (cut_facts_tac prems 1); | 
| 3457 | 60 | by (dtac DomainD 1); | 
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 61 | by (dres_inst_tac [("f","Rep_quot")] arg_cong 1);
 | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 62 | by (asm_full_simp_tac(HOL_ss addsimps [Abs_quot_inverse,per_class_rep_quot])1); | 
| 3457 | 63 | by (dres_inst_tac [("c","x")] equalityCE 1);by (assume_tac 3);
 | 
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 64 | by (safe_tac set_cs); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 65 | qed "per_class_eqE"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 66 | |
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 67 | goal thy "<[(x::'a::er)]>=<[y]>=x===y"; | 
| 3457 | 68 | by (rtac iffI 1); | 
| 69 | by (etac er_class_eqE 1); | |
| 70 | by (etac per_class_eqI 1); | |
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 71 | qed "er_class_eq"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 72 | |
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 73 | val prems = goal thy "x:D==><[x]>=<[y]>=x===y"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 74 | by (cut_facts_tac prems 1); | 
| 3457 | 75 | by (rtac iffI 1); | 
| 76 | by (etac per_class_eqE 1);by (assume_tac 1); | |
| 77 | by (etac per_class_eqI 1); | |
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 78 | qed "per_class_eq"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 79 | |
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 80 | (* inequality *) | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 81 | val prems = goal thy "[|x:D;~x===y|]==><[x]>~=<[y]>"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 82 | by (cut_facts_tac prems 1); | 
| 3457 | 83 | by (rtac notI 1); | 
| 84 | by (dtac per_class_eqE 1); | |
| 85 | by (Auto_tac()); | |
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 86 | qed "per_class_neqI"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 87 | |
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 88 | val prems = goal thy "~(x::'a::er)===y==><[x]>~=<[y]>"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 89 | by (cut_facts_tac prems 1); | 
| 3457 | 90 | by (rtac notI 1); | 
| 91 | by (dtac er_class_eqE 1); | |
| 92 | by (Auto_tac()); | |
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 93 | qed "er_class_neqI"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 94 | |
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 95 | val prems = goal thy "<[x]>~=<[y]>==>~x===y"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 96 | by (cut_facts_tac prems 1); | 
| 3457 | 97 | by (rtac notI 1); | 
| 98 | by (etac notE 1); | |
| 99 | by (etac per_class_eqI 1); | |
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 100 | qed "per_class_neqE"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 101 | |
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 102 | val prems = goal thy "x:D==><[x]>~=<[y]>=(~x===y)"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 103 | by (cut_facts_tac prems 1); | 
| 3457 | 104 | by (rtac iffI 1); | 
| 105 | by (etac per_class_neqE 1); | |
| 106 | by (etac per_class_neqI 1);by (assume_tac 1); | |
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 107 | qed "per_class_not"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 108 | |
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 109 | goal thy "<[(x::'a::er)]>~=<[y]>=(~x===y)"; | 
| 3457 | 110 | by (rtac iffI 1); | 
| 111 | by (etac per_class_neqE 1); | |
| 112 | by (etac er_class_neqI 1); | |
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 113 | qed "er_class_not"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 114 | |
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 115 | (* exhaustiveness and induction *) | 
| 3842 | 116 | goalw thy [peclass_def] "? s. x=<[s]>"; | 
| 3457 | 117 | by (rtac all_q 1); | 
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 118 | by (strip_tac 1); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 119 | by (asm_full_simp_tac (HOL_ss addsimps [mem_Collect_eq,quot_def]) 1); | 
| 3457 | 120 | by (etac exE 1); | 
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 121 | by (res_inst_tac [("x","r")] exI 1);
 | 
| 3457 | 122 | by (rtac quot_eq 1); | 
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 123 | by (subgoal_tac "s:quot" 1); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 124 | by (asm_simp_tac(HOL_ss addsimps [Abs_quot_inverse,per_class_rep_quot])1); | 
| 3457 | 125 | by (rtac set_ext 1); | 
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 126 | by (fast_tac set_cs 1); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 127 | by (asm_full_simp_tac (HOL_ss addsimps [mem_Collect_eq,quot_def]) 1); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 128 | by (fast_tac set_cs 1); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 129 | qed "per_class_exh"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 130 | |
| 3842 | 131 | val prems = goal thy "!x. P<[x]> ==> P s"; | 
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 132 | by (cut_facts_tac (prems@[ | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 133 | 	read_instantiate[("x","s::'a::per quot")] per_class_exh]) 1);
 | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 134 | by (fast_tac set_cs 1); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 135 | qed "per_class_all"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 136 | |
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 137 | (* theorems for any_in *) | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 138 | goalw thy [any_in_def,peclass_def] "any_in<[(s::'a::er)]>===s"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 139 | fr selectI2; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 140 | fr refl; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 141 | by (fold_goals_tac [peclass_def]); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 142 | fe er_class_eqE; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 143 | qed "er_any_in_class"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 144 | |
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 145 | val prems = goalw thy [any_in_def,peclass_def] "s:D==>any_in<[s]>===s"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 146 | by (cut_facts_tac prems 1); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 147 | fr selectI2; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 148 | fr refl; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 149 | by (fold_goals_tac [peclass_def]); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 150 | fr per_sym; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 151 | fe per_class_eqE; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 152 | fe sym; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 153 | qed "per_any_in_class"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 154 | |
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 155 | val prems = goal thy "<[any_in (s::'a::er quot)]> = s"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 156 | by (cut_facts_tac prems 1); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 157 | fr per_class_all; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 158 | fr allI; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 159 | fr (er_any_in_class RS per_class_eqI); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 160 | qed "er_class_any_in"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 161 | |
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 162 | (* equivalent theorem for per would need !x.x:D *) | 
| 3842 | 163 | val prems = goal thy "!x::'a::per. x:D==><[any_in (q::'a::per quot)]> = q"; | 
| 2905 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 164 | by (cut_facts_tac prems 1); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 165 | fr per_class_all; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 166 | fr allI; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 167 | fr (per_any_in_class RS per_class_eqI); | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 168 | fe spec; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 169 | qed "per_class_any_in"; | 
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 170 | |
| 
9a4f353107da
(higher-order) quotient constructor quot, based on PER
 slotosch parents: diff
changeset | 171 | (* is like theorem for class er *) |