| author | paulson | 
| Thu, 29 Jun 2000 12:15:08 +0200 | |
| changeset 9187 | 68ecc04785f1 | 
| parent 9169 | 85a47aa21f74 | 
| child 9245 | 428385c4bc50 | 
| permissions | -rw-r--r-- | 
| 2278 | 1 | (* Title: HOLCF/Up1.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Franz Regensburger | |
| 4 | Copyright 1993 Technische Universitaet Muenchen | |
| 5 | *) | |
| 6 | ||
| 9169 | 7 | Goal "Rep_Up (Abs_Up y) = y"; | 
| 8 | by (simp_tac (simpset() addsimps [Up_def,Abs_Up_inverse]) 1); | |
| 9 | qed "Abs_Up_inverse2"; | |
| 2640 | 10 | |
| 11 | qed_goalw "Exh_Up" thy [Iup_def ] | |
| 12 | "z = Abs_Up(Inl ()) | (? x. z = Iup x)" | |
| 2278 | 13 | (fn prems => | 
| 14 | [ | |
| 15 | (rtac (Rep_Up_inverse RS subst) 1), | |
| 2640 | 16 |         (res_inst_tac [("s","Rep_Up z")] sumE 1),
 | 
| 2278 | 17 | (rtac disjI1 1), | 
| 18 |         (res_inst_tac [("f","Abs_Up")] arg_cong 1),
 | |
| 2640 | 19 | (rtac (unit_eq RS subst) 1), | 
| 2278 | 20 | (atac 1), | 
| 21 | (rtac disjI2 1), | |
| 22 | (rtac exI 1), | |
| 23 |         (res_inst_tac [("f","Abs_Up")] arg_cong 1),
 | |
| 24 | (atac 1) | |
| 25 | ]); | |
| 26 | ||
| 9169 | 27 | Goal "inj(Abs_Up)"; | 
| 28 | by (rtac inj_inverseI 1); | |
| 29 | by (rtac Abs_Up_inverse2 1); | |
| 30 | qed "inj_Abs_Up"; | |
| 2278 | 31 | |
| 9169 | 32 | Goal "inj(Rep_Up)"; | 
| 33 | by (rtac inj_inverseI 1); | |
| 34 | by (rtac Rep_Up_inverse 1); | |
| 35 | qed "inj_Rep_Up"; | |
| 2278 | 36 | |
| 2640 | 37 | qed_goalw "inject_Iup" thy [Iup_def] "Iup x=Iup y ==> x=y" | 
| 2278 | 38 | (fn prems => | 
| 39 | [ | |
| 40 | (cut_facts_tac prems 1), | |
| 41 | (rtac (inj_Inr RS injD) 1), | |
| 42 | (rtac (inj_Abs_Up RS injD) 1), | |
| 43 | (atac 1) | |
| 44 | ]); | |
| 45 | ||
| 9169 | 46 | AddSDs [inject_Iup]; | 
| 47 | ||
| 2640 | 48 | qed_goalw "defined_Iup" thy [Iup_def] "Iup x~=Abs_Up(Inl ())" | 
| 2278 | 49 | (fn prems => | 
| 50 | [ | |
| 51 | (rtac notI 1), | |
| 52 | (rtac notE 1), | |
| 53 | (rtac Inl_not_Inr 1), | |
| 54 | (rtac sym 1), | |
| 55 | (etac (inj_Abs_Up RS injD) 1) | |
| 56 | ]); | |
| 57 | ||
| 58 | ||
| 9169 | 59 | val prems = Goal "[| p=Abs_Up(Inl ()) ==> Q; !!x. p=Iup(x)==>Q|] ==>Q"; | 
| 60 | by (rtac (Exh_Up RS disjE) 1); | |
| 61 | by (eresolve_tac prems 1); | |
| 62 | by (etac exE 1); | |
| 63 | by (eresolve_tac prems 1); | |
| 64 | qed "upE"; | |
| 2278 | 65 | |
| 2640 | 66 | qed_goalw "Ifup1" thy [Ifup_def] | 
| 67 | "Ifup(f)(Abs_Up(Inl ()))=UU" | |
| 2278 | 68 | (fn prems => | 
| 69 | [ | |
| 2640 | 70 | (stac Abs_Up_inverse2 1), | 
| 2278 | 71 | (stac sum_case_Inl 1), | 
| 72 | (rtac refl 1) | |
| 73 | ]); | |
| 74 | ||
| 2640 | 75 | qed_goalw "Ifup2" thy [Ifup_def,Iup_def] | 
| 2278 | 76 | "Ifup(f)(Iup(x))=f`x" | 
| 77 | (fn prems => | |
| 78 | [ | |
| 2640 | 79 | (stac Abs_Up_inverse2 1), | 
| 2278 | 80 | (stac sum_case_Inr 1), | 
| 81 | (rtac refl 1) | |
| 82 | ]); | |
| 83 | ||
| 8161 | 84 | val Up0_ss = (simpset_of Cfun3.thy) delsimps [range_composition] | 
| 85 | addsimps [Ifup1,Ifup2]; | |
| 2278 | 86 | |
| 9169 | 87 | Addsimps [Ifup1,Ifup2]; | 
| 88 | ||
| 2640 | 89 | qed_goalw "less_up1a" thy [less_up_def] | 
| 3323 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
2640diff
changeset | 90 | "Abs_Up(Inl ())<< z" | 
| 2278 | 91 | (fn prems => | 
| 92 | [ | |
| 2640 | 93 | (stac Abs_Up_inverse2 1), | 
| 2278 | 94 | (stac sum_case_Inl 1), | 
| 95 | (rtac TrueI 1) | |
| 96 | ]); | |
| 97 | ||
| 2640 | 98 | qed_goalw "less_up1b" thy [Iup_def,less_up_def] | 
| 3323 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 slotosch parents: 
2640diff
changeset | 99 | "~(Iup x) << (Abs_Up(Inl ()))" | 
| 2278 | 100 | (fn prems => | 
| 101 | [ | |
| 102 | (rtac notI 1), | |
| 103 | (rtac iffD1 1), | |
| 104 | (atac 2), | |
| 2640 | 105 | (stac Abs_Up_inverse2 1), | 
| 106 | (stac Abs_Up_inverse2 1), | |
| 2278 | 107 | (stac sum_case_Inr 1), | 
| 108 | (stac sum_case_Inl 1), | |
| 109 | (rtac refl 1) | |
| 110 | ]); | |
| 111 | ||
| 2640 | 112 | qed_goalw "less_up1c" thy [Iup_def,less_up_def] | 
| 9169 | 113 | "(Iup x) << (Iup y)=(x<<y)" | 
| 2278 | 114 | (fn prems => | 
| 115 | [ | |
| 2640 | 116 | (stac Abs_Up_inverse2 1), | 
| 117 | (stac Abs_Up_inverse2 1), | |
| 2278 | 118 | (stac sum_case_Inr 1), | 
| 119 | (stac sum_case_Inr 1), | |
| 120 | (rtac refl 1) | |
| 121 | ]); | |
| 122 | ||
| 9169 | 123 | AddIffs [less_up1a, less_up1b, less_up1c]; | 
| 2278 | 124 | |
| 9169 | 125 | Goal "(p::'a u) << p"; | 
| 126 | by (res_inst_tac [("p","p")] upE 1);
 | |
| 127 | by Auto_tac; | |
| 128 | qed "refl_less_up"; | |
| 2278 | 129 | |
| 9169 | 130 | Goal "[|(p1::'a u) << p2;p2 << p1|] ==> p1=p2"; | 
| 131 | by (res_inst_tac [("p","p1")] upE 1);
 | |
| 132 | by (hyp_subst_tac 1); | |
| 133 | by (res_inst_tac [("p","p2")] upE 1);
 | |
| 134 | by (etac sym 1); | |
| 135 | by (hyp_subst_tac 1); | |
| 136 | by (res_inst_tac [("P","(Iup x) << (Abs_Up(Inl ()))")] notE 1);
 | |
| 137 | by (rtac less_up1b 1); | |
| 138 | by (atac 1); | |
| 139 | by (hyp_subst_tac 1); | |
| 140 | by (res_inst_tac [("p","p2")] upE 1);
 | |
| 141 | by (hyp_subst_tac 1); | |
| 142 | by (res_inst_tac [("P","(Iup x) << (Abs_Up(Inl ()))")] notE 1);
 | |
| 143 | by (rtac less_up1b 1); | |
| 144 | by (atac 1); | |
| 145 | by (blast_tac (claset() addIs [arg_cong, antisym_less, less_up1c RS iffD1]) 1); | |
| 146 | qed "antisym_less_up"; | |
| 2278 | 147 | |
| 9169 | 148 | Goal "[|(p1::'a u) << p2;p2 << p3|] ==> p1 << p3"; | 
| 149 | by (res_inst_tac [("p","p1")] upE 1);
 | |
| 150 | by (hyp_subst_tac 1); | |
| 151 | by (rtac less_up1a 1); | |
| 152 | by (hyp_subst_tac 1); | |
| 153 | by (res_inst_tac [("p","p2")] upE 1);
 | |
| 154 | by (hyp_subst_tac 1); | |
| 155 | by (rtac notE 1); | |
| 156 | by (rtac less_up1b 1); | |
| 157 | by (atac 1); | |
| 158 | by (res_inst_tac [("p","p3")] upE 1);
 | |
| 159 | by Auto_tac; | |
| 160 | by (blast_tac (claset() addIs [trans_less]) 1); | |
| 161 | qed "trans_less_up"; | |
| 2278 | 162 |