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