X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fng_assembly%2Ffreescale%2Ftranslation.ma;h=0996948942bad7c8c8cd7b4809456eaa72acc96b;hb=ec07ff398325533d848da92e9dc69852d24b78a5;hp=0474e6e9bf88bf508d912f50e792202c477f1e22;hpb=c70ecb50a457d251ef1cd61960a641d491febed7;p=helm.git diff --git a/helm/software/matita/contribs/ng_assembly/freescale/translation.ma b/helm/software/matita/contribs/ng_assembly/freescale/translation.ma index 0474e6e9b..099694894 100755 --- a/helm/software/matita/contribs/ng_assembly/freescale/translation.ma +++ b/helm/software/matita/contribs/ng_assembly/freescale/translation.ma @@ -15,8 +15,8 @@ (* ********************************************************************** *) (* Progetto FreeScale *) (* *) -(* Sviluppato da: Cosimo Oliboni, oliboni@cs.unibo.it *) -(* Cosimo Oliboni, oliboni@cs.unibo.it *) +(* Sviluppato da: Ing. Cosimo Oliboni, oliboni@cs.unibo.it *) +(* Ultima modifica: 05/08/2009 *) (* *) (* ********************************************************************** *) @@ -76,13 +76,11 @@ full_info_of_word16_aux m borw (opcode_table m). ninductive t_byte8 (m:mcu_type) : Type ≝ TByte : byte8 → t_byte8 m. -nlemma tbyte8_destruct : ∀m,b1,b2.TByte m b1 = TByte m b2 → b1 = b2. - #m; #b1; #b2; #H; - nchange with (match TByte m b2 with [ TByte a ⇒ b1 = a ]); - nrewrite < H; - nnormalize; - napply refl_eq. -nqed. +ndefinition eq_tbyte8 ≝ +λm.λtb1,tb2:t_byte8 m. + match tb1 with + [ TByte b1 ⇒ match tb2 with + [ TByte b2 ⇒ eq_b8 b1 b2 ]]. (* introduzione di un tipo dipendente (dalla modalita') per gli argomenti *) ninductive MA_check : instr_mode → Type ≝