| author | nipkow | 
| Tue, 22 Feb 2005 10:54:30 +0100 | |
| changeset 15543 | 0024472afce7 | 
| parent 14981 | e73f8140af78 | 
| permissions | -rw-r--r-- | 
| 2278 | 1 | (* Title: HOLCF/Up1.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Franz Regensburger | |
| 9245 | 4 | |
| 12030 | 5 | Lifting. | 
| 2278 | 6 | *) | 
| 7 | ||
| 9169 | 8 | Goal "Rep_Up (Abs_Up y) = y"; | 
| 9 | by (simp_tac (simpset() addsimps [Up_def,Abs_Up_inverse]) 1); | |
| 10 | qed "Abs_Up_inverse2"; | |
| 2640 | 11 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 12 | Goalw [Iup_def] "z = Abs_Up(Inl ()) | (? x. z = Iup x)"; | 
| 9245 | 13 | by (rtac (Rep_Up_inverse RS subst) 1); | 
| 14 | by (res_inst_tac [("s","Rep_Up z")] sumE 1);
 | |
| 15 | by (rtac disjI1 1); | |
| 16 | by (res_inst_tac [("f","Abs_Up")] arg_cong 1);
 | |
| 17 | by (rtac (unit_eq RS subst) 1); | |
| 18 | by (atac 1); | |
| 19 | by (rtac disjI2 1); | |
| 20 | by (rtac exI 1); | |
| 21 | by (res_inst_tac [("f","Abs_Up")] arg_cong 1);
 | |
| 22 | by (atac 1); | |
| 23 | qed "Exh_Up"; | |
| 2278 | 24 | |
| 9169 | 25 | Goal "inj(Abs_Up)"; | 
| 26 | by (rtac inj_inverseI 1); | |
| 27 | by (rtac Abs_Up_inverse2 1); | |
| 28 | qed "inj_Abs_Up"; | |
| 2278 | 29 | |
| 9169 | 30 | Goal "inj(Rep_Up)"; | 
| 31 | by (rtac inj_inverseI 1); | |
| 32 | by (rtac Rep_Up_inverse 1); | |
| 33 | qed "inj_Rep_Up"; | |
| 2278 | 34 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 35 | Goalw [Iup_def] "Iup x=Iup y ==> x=y"; | 
| 9245 | 36 | by (rtac (inj_Inr RS injD) 1); | 
| 37 | by (rtac (inj_Abs_Up RS injD) 1); | |
| 38 | by (atac 1); | |
| 39 | qed "inject_Iup"; | |
| 2278 | 40 | |
| 9169 | 41 | AddSDs [inject_Iup]; | 
| 42 | ||
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 43 | Goalw [Iup_def] "Iup x~=Abs_Up(Inl ())"; | 
| 9245 | 44 | by (rtac notI 1); | 
| 45 | by (rtac notE 1); | |
| 46 | by (rtac Inl_not_Inr 1); | |
| 47 | by (rtac sym 1); | |
| 48 | by (etac (inj_Abs_Up RS injD) 1); | |
| 49 | qed "defined_Iup"; | |
| 2278 | 50 | |
| 51 | ||
| 9169 | 52 | val prems = Goal "[| p=Abs_Up(Inl ()) ==> Q; !!x. p=Iup(x)==>Q|] ==>Q"; | 
| 53 | by (rtac (Exh_Up RS disjE) 1); | |
| 54 | by (eresolve_tac prems 1); | |
| 55 | by (etac exE 1); | |
| 56 | by (eresolve_tac prems 1); | |
| 57 | qed "upE"; | |
| 2278 | 58 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 59 | Goalw [Ifup_def] "Ifup(f)(Abs_Up(Inl ()))=UU"; | 
| 9245 | 60 | by (stac Abs_Up_inverse2 1); | 
| 61 | by (stac sum_case_Inl 1); | |
| 62 | by (rtac refl 1); | |
| 63 | qed "Ifup1"; | |
| 2278 | 64 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 65 | Goalw [Ifup_def,Iup_def] | 
| 10834 | 66 | "Ifup(f)(Iup(x))=f$x"; | 
| 9245 | 67 | by (stac Abs_Up_inverse2 1); | 
| 68 | by (stac sum_case_Inr 1); | |
| 69 | by (rtac refl 1); | |
| 70 | qed "Ifup2"; | |
| 2278 | 71 | |
| 8161 | 72 | val Up0_ss = (simpset_of Cfun3.thy) delsimps [range_composition] | 
| 73 | addsimps [Ifup1,Ifup2]; | |
| 2278 | 74 | |
| 9169 | 75 | Addsimps [Ifup1,Ifup2]; | 
| 76 | ||
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 77 | Goalw [less_up_def] | 
| 9245 | 78 | "Abs_Up(Inl ())<< z"; | 
| 79 | by (stac Abs_Up_inverse2 1); | |
| 80 | by (stac sum_case_Inl 1); | |
| 81 | by (rtac TrueI 1); | |
| 82 | qed "less_up1a"; | |
| 2278 | 83 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 84 | Goalw [Iup_def,less_up_def] | 
| 9245 | 85 | "~(Iup x) << (Abs_Up(Inl ()))"; | 
| 86 | by (rtac notI 1); | |
| 87 | by (rtac iffD1 1); | |
| 88 | by (atac 2); | |
| 89 | by (stac Abs_Up_inverse2 1); | |
| 90 | by (stac Abs_Up_inverse2 1); | |
| 91 | by (stac sum_case_Inr 1); | |
| 92 | by (stac sum_case_Inl 1); | |
| 93 | by (rtac refl 1); | |
| 94 | qed "less_up1b"; | |
| 2278 | 95 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 96 | Goalw [Iup_def,less_up_def] | 
| 9245 | 97 | "(Iup x) << (Iup y)=(x<<y)"; | 
| 98 | by (stac Abs_Up_inverse2 1); | |
| 99 | by (stac Abs_Up_inverse2 1); | |
| 100 | by (stac sum_case_Inr 1); | |
| 101 | by (stac sum_case_Inr 1); | |
| 102 | by (rtac refl 1); | |
| 103 | qed "less_up1c"; | |
| 2278 | 104 | |
| 9169 | 105 | AddIffs [less_up1a, less_up1b, less_up1c]; | 
| 2278 | 106 | |
| 9169 | 107 | Goal "(p::'a u) << p"; | 
| 108 | by (res_inst_tac [("p","p")] upE 1);
 | |
| 109 | by Auto_tac; | |
| 110 | qed "refl_less_up"; | |
| 2278 | 111 | |
| 9169 | 112 | Goal "[|(p1::'a u) << p2;p2 << p1|] ==> p1=p2"; | 
| 113 | by (res_inst_tac [("p","p1")] upE 1);
 | |
| 114 | by (hyp_subst_tac 1); | |
| 115 | by (res_inst_tac [("p","p2")] upE 1);
 | |
| 116 | by (etac sym 1); | |
| 117 | by (hyp_subst_tac 1); | |
| 118 | by (res_inst_tac [("P","(Iup x) << (Abs_Up(Inl ()))")] notE 1);
 | |
| 119 | by (rtac less_up1b 1); | |
| 120 | by (atac 1); | |
| 121 | by (hyp_subst_tac 1); | |
| 122 | by (res_inst_tac [("p","p2")] upE 1);
 | |
| 123 | by (hyp_subst_tac 1); | |
| 124 | by (res_inst_tac [("P","(Iup x) << (Abs_Up(Inl ()))")] notE 1);
 | |
| 125 | by (rtac less_up1b 1); | |
| 126 | by (atac 1); | |
| 127 | by (blast_tac (claset() addIs [arg_cong, antisym_less, less_up1c RS iffD1]) 1); | |
| 128 | qed "antisym_less_up"; | |
| 2278 | 129 | |
| 9169 | 130 | Goal "[|(p1::'a u) << p2;p2 << p3|] ==> p1 << p3"; | 
| 131 | by (res_inst_tac [("p","p1")] upE 1);
 | |
| 132 | by (hyp_subst_tac 1); | |
| 133 | by (rtac less_up1a 1); | |
| 134 | by (hyp_subst_tac 1); | |
| 135 | by (res_inst_tac [("p","p2")] upE 1);
 | |
| 136 | by (hyp_subst_tac 1); | |
| 137 | by (rtac notE 1); | |
| 138 | by (rtac less_up1b 1); | |
| 139 | by (atac 1); | |
| 140 | by (res_inst_tac [("p","p3")] upE 1);
 | |
| 141 | by Auto_tac; | |
| 142 | by (blast_tac (claset() addIs [trans_less]) 1); | |
| 143 | qed "trans_less_up"; | |
| 2278 | 144 |