| 10341 |      1 | (* ID:         $Id$ *)
 | 
| 16417 |      2 | theory Functions imports Main begin
 | 
| 10294 |      3 | 
 | 
|  |      4 | ML "Pretty.setmargin 64"
 | 
|  |      5 | 
 | 
|  |      6 | 
 | 
|  |      7 | text{*
 | 
|  |      8 | @{thm[display] id_def[no_vars]}
 | 
|  |      9 | \rulename{id_def}
 | 
|  |     10 | 
 | 
|  |     11 | @{thm[display] o_def[no_vars]}
 | 
|  |     12 | \rulename{o_def}
 | 
|  |     13 | 
 | 
|  |     14 | @{thm[display] o_assoc[no_vars]}
 | 
|  |     15 | \rulename{o_assoc}
 | 
|  |     16 | *}
 | 
|  |     17 | 
 | 
|  |     18 | text{*
 | 
|  |     19 | @{thm[display] fun_upd_apply[no_vars]}
 | 
|  |     20 | \rulename{fun_upd_apply}
 | 
|  |     21 | 
 | 
|  |     22 | @{thm[display] fun_upd_upd[no_vars]}
 | 
|  |     23 | \rulename{fun_upd_upd}
 | 
|  |     24 | *}
 | 
|  |     25 | 
 | 
|  |     26 | 
 | 
|  |     27 | text{*
 | 
|  |     28 | definitions of injective, surjective, bijective
 | 
|  |     29 | 
 | 
|  |     30 | @{thm[display] inj_on_def[no_vars]}
 | 
|  |     31 | \rulename{inj_on_def}
 | 
|  |     32 | 
 | 
|  |     33 | @{thm[display] surj_def[no_vars]}
 | 
|  |     34 | \rulename{surj_def}
 | 
|  |     35 | 
 | 
|  |     36 | @{thm[display] bij_def[no_vars]}
 | 
|  |     37 | \rulename{bij_def}
 | 
|  |     38 | *}
 | 
|  |     39 | 
 | 
|  |     40 | 
 | 
|  |     41 | 
 | 
|  |     42 | text{*
 | 
|  |     43 | possibly interesting theorems about inv
 | 
|  |     44 | *}
 | 
|  |     45 | 
 | 
|  |     46 | text{*
 | 
|  |     47 | @{thm[display] inv_f_f[no_vars]}
 | 
|  |     48 | \rulename{inv_f_f}
 | 
|  |     49 | 
 | 
|  |     50 | @{thm[display] inj_imp_surj_inv[no_vars]}
 | 
|  |     51 | \rulename{inj_imp_surj_inv}
 | 
|  |     52 | 
 | 
|  |     53 | @{thm[display] surj_imp_inj_inv[no_vars]}
 | 
|  |     54 | \rulename{surj_imp_inj_inv}
 | 
|  |     55 | 
 | 
|  |     56 | @{thm[display] surj_f_inv_f[no_vars]}
 | 
|  |     57 | \rulename{surj_f_inv_f}
 | 
|  |     58 | 
 | 
|  |     59 | @{thm[display] bij_imp_bij_inv[no_vars]}
 | 
|  |     60 | \rulename{bij_imp_bij_inv}
 | 
|  |     61 | 
 | 
|  |     62 | @{thm[display] inv_inv_eq[no_vars]}
 | 
|  |     63 | \rulename{inv_inv_eq}
 | 
|  |     64 | 
 | 
|  |     65 | @{thm[display] o_inv_distrib[no_vars]}
 | 
|  |     66 | \rulename{o_inv_distrib}
 | 
|  |     67 | *}
 | 
|  |     68 | 
 | 
|  |     69 | 
 | 
|  |     70 | 
 | 
|  |     71 | text{*
 | 
|  |     72 | small sample proof
 | 
|  |     73 | 
 | 
|  |     74 | @{thm[display] ext[no_vars]}
 | 
|  |     75 | \rulename{ext}
 | 
|  |     76 | 
 | 
|  |     77 | @{thm[display] expand_fun_eq[no_vars]}
 | 
|  |     78 | \rulename{expand_fun_eq}
 | 
|  |     79 | *}
 | 
|  |     80 | 
 | 
|  |     81 | lemma "inj f \<Longrightarrow> (f o g = f o h) = (g = h)";
 | 
| 10983 |     82 |   apply (simp add: expand_fun_eq inj_on_def)
 | 
| 10294 |     83 |   apply (auto)
 | 
|  |     84 |   done
 | 
|  |     85 | 
 | 
|  |     86 | text{*
 | 
|  |     87 | \begin{isabelle}
 | 
|  |     88 | inj\ f\ \isasymLongrightarrow \ (f\ \isasymcirc \ g\ =\ f\ \isasymcirc \ h)\ =\ (g\ =\ h)\isanewline
 | 
|  |     89 | \ 1.\ \isasymforall x\ y.\ f\ x\ =\ f\ y\ \isasymlongrightarrow \ x\ =\ y\ \isasymLongrightarrow \isanewline
 | 
|  |     90 | \ \ \ \ (\isasymforall x.\ f\ (g\ x)\ =\ f\ (h\ x))\ =\ (\isasymforall x.\ g\ x\ =\ h\ x)
 | 
|  |     91 | \end{isabelle}
 | 
|  |     92 | *}
 | 
|  |     93 |  
 | 
|  |     94 | 
 | 
|  |     95 | text{*image, inverse image*}
 | 
|  |     96 | 
 | 
|  |     97 | text{*
 | 
|  |     98 | @{thm[display] image_def[no_vars]}
 | 
|  |     99 | \rulename{image_def}
 | 
|  |    100 | *}
 | 
|  |    101 | 
 | 
|  |    102 | text{*
 | 
|  |    103 | @{thm[display] image_Un[no_vars]}
 | 
|  |    104 | \rulename{image_Un}
 | 
|  |    105 | *}
 | 
|  |    106 | 
 | 
|  |    107 | text{*
 | 
|  |    108 | @{thm[display] image_compose[no_vars]}
 | 
|  |    109 | \rulename{image_compose}
 | 
|  |    110 | 
 | 
|  |    111 | @{thm[display] image_Int[no_vars]}
 | 
|  |    112 | \rulename{image_Int}
 | 
|  |    113 | 
 | 
|  |    114 | @{thm[display] bij_image_Compl_eq[no_vars]}
 | 
|  |    115 | \rulename{bij_image_Compl_eq}
 | 
|  |    116 | *}
 | 
|  |    117 | 
 | 
|  |    118 | 
 | 
|  |    119 | text{*
 | 
|  |    120 | illustrates Union as well as image
 | 
|  |    121 | *}
 | 
| 10849 |    122 | 
 | 
| 10839 |    123 | lemma "f`A \<union> g`A = (\<Union>x\<in>A. {f x, g x})"
 | 
| 10849 |    124 | by blast
 | 
| 10294 |    125 | 
 | 
| 10839 |    126 | lemma "f ` {(x,y). P x y} = {f(x,y) | x y. P x y}"
 | 
| 10849 |    127 | by blast
 | 
| 10294 |    128 | 
 | 
|  |    129 | text{*actually a macro!*}
 | 
|  |    130 | 
 | 
| 10839 |    131 | lemma "range f = f`UNIV"
 | 
| 10849 |    132 | by blast
 | 
| 10294 |    133 | 
 | 
|  |    134 | 
 | 
|  |    135 | text{*
 | 
|  |    136 | inverse image
 | 
|  |    137 | *}
 | 
|  |    138 | 
 | 
|  |    139 | text{*
 | 
|  |    140 | @{thm[display] vimage_def[no_vars]}
 | 
|  |    141 | \rulename{vimage_def}
 | 
|  |    142 | 
 | 
|  |    143 | @{thm[display] vimage_Compl[no_vars]}
 | 
|  |    144 | \rulename{vimage_Compl}
 | 
|  |    145 | *}
 | 
|  |    146 | 
 | 
|  |    147 | 
 | 
|  |    148 | end
 |