]> matita.cs.unibo.it Git - helm.git/commitdiff
Old tiny freescale experiment get rid of.
authorClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Mon, 21 Jan 2008 17:29:48 +0000 (17:29 +0000)
committerClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Mon, 21 Jan 2008 17:29:48 +0000 (17:29 +0000)
New shiny freescale formalization by Cosimo Oliboni <oliboni@cs.unibo.it>,
described in his Laurea Thesis.

36 files changed:
helm/software/matita/library/assembly/byte.ma [deleted file]
helm/software/matita/library/assembly/exadecimal.ma [deleted file]
helm/software/matita/library/assembly/extra.ma [deleted file]
helm/software/matita/library/assembly/test.ma [deleted file]
helm/software/matita/library/assembly/vm.ma [deleted file]
helm/software/matita/library/depends
helm/software/matita/library/freescale/aux_bases.ma [new file with mode: 0644]
helm/software/matita/library/freescale/byte8.ma [new file with mode: 0644]
helm/software/matita/library/freescale/doc/aurei.txt [new file with mode: 0644]
helm/software/matita/library/freescale/doc/daa.txt [new file with mode: 0644]
helm/software/matita/library/freescale/doc/freescale.txt [new file with mode: 0644]
helm/software/matita/library/freescale/doc/ordinamento.txt [new file with mode: 0644]
helm/software/matita/library/freescale/doc/ordine_compilazione.txt [new file with mode: 0644]
helm/software/matita/library/freescale/doc/reverse.txt [new file with mode: 0644]
helm/software/matita/library/freescale/exadecim.ma [new file with mode: 0644]
helm/software/matita/library/freescale/extra.ma [new file with mode: 0644]
helm/software/matita/library/freescale/load_write.ma [new file with mode: 0644]
helm/software/matita/library/freescale/medium_tests.ma [new file with mode: 0644]
helm/software/matita/library/freescale/medium_tests_tools.ma [new file with mode: 0644]
helm/software/matita/library/freescale/memory_abs.ma [new file with mode: 0644]
helm/software/matita/library/freescale/memory_bits.ma [new file with mode: 0644]
helm/software/matita/library/freescale/memory_func.ma [new file with mode: 0644]
helm/software/matita/library/freescale/memory_struct.ma [new file with mode: 0644]
helm/software/matita/library/freescale/memory_trees.ma [new file with mode: 0644]
helm/software/matita/library/freescale/micro_tests.ma [new file with mode: 0644]
helm/software/matita/library/freescale/model.ma [new file with mode: 0644]
helm/software/matita/library/freescale/multivm.ma [new file with mode: 0644]
helm/software/matita/library/freescale/opcode.ma [new file with mode: 0644]
helm/software/matita/library/freescale/status.ma [new file with mode: 0644]
helm/software/matita/library/freescale/table_HC05.ma [new file with mode: 0644]
helm/software/matita/library/freescale/table_HC08.ma [new file with mode: 0644]
helm/software/matita/library/freescale/table_HCS08.ma [new file with mode: 0644]
helm/software/matita/library/freescale/table_RS08.ma [new file with mode: 0644]
helm/software/matita/library/freescale/tests.old [new file with mode: 0644]
helm/software/matita/library/freescale/translation.ma [new file with mode: 0644]
helm/software/matita/library/freescale/word16.ma [new file with mode: 0644]

diff --git a/helm/software/matita/library/assembly/byte.ma b/helm/software/matita/library/assembly/byte.ma
deleted file mode 100644 (file)
index 2bf233e..0000000
+++ /dev/null
@@ -1,334 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||M||                                                             *)
-(*      ||A||       A project by Andrea Asperti                           *)
-(*      ||T||                                                             *)
-(*      ||I||       Developers:                                           *)
-(*      ||T||         The HELM team.                                      *)
-(*      ||A||         http://helm.cs.unibo.it                             *)
-(*      \   /                                                             *)
-(*       \ /        This file is distributed under the terms of the       *)
-(*        v         GNU General Public License Version 2                  *)
-(*                                                                        *)
-(**************************************************************************)
-
-set "baseuri" "cic:/matita/assembly/byte".
-
-include "assembly/exadecimal.ma".
-
-record byte : Type ≝ {
- bh: exadecimal;
- bl: exadecimal
-}.
-
-notation "〈 x, y 〉" non associative with precedence 80 for @{ 'mk_byte $x $y }.
-interpretation "mk_byte" 'mk_byte x y = 
- (cic:/matita/assembly/byte/byte.ind#xpointer(1/1/1) x y).
-
-definition eqbyte ≝
- λb,b'. eqex (bh b) (bh b') ∧ eqex (bl b) (bl b').
-
-definition plusbyte ≝
- λb1,b2,c.
-  match plusex (bl b1) (bl b2) c with
-   [ couple l c' ⇒
-      match plusex (bh b1) (bh b2) c' with
-       [ couple h c'' ⇒ couple ? ? (mk_byte h l) c'' ]].
-
-definition nat_of_byte ≝ λb:byte. 16*(bh b) + (bl b).
-
-coercion cic:/matita/assembly/byte/nat_of_byte.con.
-
-definition byte_of_nat ≝
- λn. mk_byte (exadecimal_of_nat (n / 16)) (exadecimal_of_nat n).
-
-interpretation "byte_of_nat" 'byte_of_opcode a =
- (cic:/matita/assembly/byte/byte_of_nat.con a).
-
-lemma byte_of_nat_nat_of_byte: ∀b. byte_of_nat (nat_of_byte b) = b.
- intros;
- elim b;
- elim e;
- elim e1;
- reflexivity.
-qed.
-
-lemma lt_nat_of_byte_256: ∀b. nat_of_byte b < 256.
- intro;
- unfold nat_of_byte;
- letin H ≝ (lt_nat_of_exadecimal_16 (bh b)); clearbody H;
- letin K ≝ (lt_nat_of_exadecimal_16 (bl b)); clearbody K;
- unfold lt in H K ⊢ %;
- letin H' ≝ (le_S_S_to_le ? ? H); clearbody H'; clear H;
- letin K' ≝ (le_S_S_to_le ? ? K); clearbody K'; clear K;
- apply le_S_S;
- cut (16*bh b ≤ 16*15);
-  [ letin Hf ≝ (le_plus ? ? ? ? Hcut K'); clearbody Hf;
-    simplify in Hf:(? ? %);
-    assumption
-  | apply le_times_r. apply H'.
-  ]
-qed.
-
-lemma nat_of_byte_byte_of_nat: ∀n. nat_of_byte (byte_of_nat n) = n \mod 256.
- intro;
- letin H ≝ (lt_nat_of_byte_256 (byte_of_nat n)); clearbody H;
- rewrite < (lt_to_eq_mod ? ? H); clear H;
- unfold byte_of_nat;
- unfold nat_of_byte;
- change with ((16*(exadecimal_of_nat (n/16)) + exadecimal_of_nat n) \mod 256 = n \mod 256);
- letin H ≝ (div_mod n 16 ?); clearbody H; [ autobatch | ];
- rewrite > symmetric_times in H;
- rewrite > nat_of_exadecimal_exadecimal_of_nat in ⊢ (? ? (? (? % ?) ?) ?);
- rewrite > nat_of_exadecimal_exadecimal_of_nat in ⊢ (? ? (? (? ? %) ?) ?);
- rewrite > H in ⊢ (? ? ? (? % ?)); clear H;
- rewrite > mod_plus in ⊢ (? ? % ?);
- rewrite > mod_plus in ⊢ (? ? ? %);
- apply eq_mod_to_eq_plus_mod;
- rewrite < mod_mod in ⊢ (? ? ? %); [ | autobatch];
- rewrite < mod_mod in ⊢ (? ? % ?); [ | autobatch];
- rewrite < (eq_mod_times_times_mod ? ? 16 256) in ⊢ (? ? (? % ?) ?); [2: reflexivity | ];
- rewrite < mod_mod in ⊢ (? ? % ?);
-  [ reflexivity
-  | autobatch
-  ].
-qed.
-
-lemma eq_nat_of_byte_n_nat_of_byte_mod_n_256:
- ∀n. byte_of_nat n = byte_of_nat (n \mod 256).
- intro;
- unfold byte_of_nat;
- apply eq_f2;
-  [ rewrite > exadecimal_of_nat_mod in ⊢ (? ? % ?);
-    rewrite > exadecimal_of_nat_mod in ⊢ (? ? ? %);
-    apply eq_f;
-    elim daemon
-  | rewrite > exadecimal_of_nat_mod;
-    rewrite > exadecimal_of_nat_mod in ⊢ (? ? ? %);
-    rewrite > divides_to_eq_mod_mod_mod;
-     [ reflexivity
-     | apply (witness ? ? 16). reflexivity.
-     ]
-  ]
-qed.
-
-lemma plusbyte_ok:
- ∀b1,b2,c.
-  match plusbyte b1 b2 c with
-   [ couple r c' ⇒ b1 + b2 + nat_of_bool c = nat_of_byte r + nat_of_bool c' * 256
-   ].
- intros; elim daemon.
- (* 
- unfold plusbyte;
- generalize in match (plusex_ok (bl b1) (bl b2) c);
- elim (plusex (bl b1) (bl b2) c);
- simplify in H ⊢ %;
- generalize in match (plusex_ok (bh b1) (bh b2) t1);
- elim (plusex (bh b1) (bh b2) t1);
- simplify in H1 ⊢ %;
- change in ⊢ (? ? ? (? (? % ?) ?)) with (16 * t2);
- unfold nat_of_byte;
- letin K ≝ (eq_f ? ? (λy.16*y) ? ? H1); clearbody K; clear H1;
- rewrite > distr_times_plus in K:(? ? ? %);
- rewrite > symmetric_times in K:(? ? ? (? ? (? ? %)));
- rewrite < associative_times in K:(? ? ? (? ? %));
- normalize in K:(? ? ? (? ? (? % ?)));
- rewrite > symmetric_times in K:(? ? ? (? ? %));
- rewrite > sym_plus in ⊢ (? ? ? (? % ?));
- rewrite > associative_plus in ⊢ (? ? ? %);
- letin K' ≝ (eq_f ? ? (plus t) ? ? K); clearbody K'; clear K;
-  apply transitive_eq; [3: apply K' | skip | ];
- clear K';
- rewrite > sym_plus in ⊢ (? ? (? (? ? %) ?) ?);
- rewrite > associative_plus in ⊢ (? ? (? % ?) ?);
- rewrite > associative_plus in ⊢ (? ? % ?);
- rewrite > associative_plus in ⊢ (? ? (? ? %) ?);
- rewrite > associative_plus in ⊢ (? ? (? ? (? ? %)) ?);
- rewrite > sym_plus in ⊢ (? ? (? ? (? ? (? ? %))) ?);
- rewrite < associative_plus in ⊢ (? ? (? ? (? ? %)) ?);
- rewrite < associative_plus in ⊢ (? ? (? ? %) ?);
- rewrite < associative_plus in ⊢ (? ? (? ? (? % ?)) ?);
- rewrite > H; clear H;
- autobatch paramodulation.
- *)
-qed.
-
-definition bpred ≝
- λb.
-  match eqex (bl b) x0 with
-   [ true ⇒ mk_byte (xpred (bh b)) (xpred (bl b))
-   | false ⇒ mk_byte (bh b) (xpred (bl b))
-   ]. 
-
-lemma plusbyte_O_x:
- ∀b. plusbyte (mk_byte x0 x0) b false = couple ? ? b false.
- intros;
- elim b;
- elim e;
- elim e1;
- reflexivity.
-qed.
-
-definition plusbytenc ≝
- λx,y.
-  match plusbyte x y false with
-   [couple res _ ⇒ res].
-
-definition plusbytec ≝
- λx,y.
-  match plusbyte x y false with
-   [couple _ c ⇒ c].
-
-lemma plusbytenc_O_x:
- ∀x. plusbytenc (mk_byte x0 x0) x = x.
- intros;
- unfold plusbytenc;
- rewrite > plusbyte_O_x;
- reflexivity.
-qed.
-
-lemma eq_nat_of_byte_mod: ∀b. nat_of_byte b = nat_of_byte b \mod 256.
- intro;
- lapply (lt_nat_of_byte_256 b);
- rewrite > (lt_to_eq_mod ? ? Hletin) in ⊢ (? ? ? %);
- reflexivity.
-qed.
-
-theorem plusbytenc_ok:
- ∀b1,b2:byte. nat_of_byte (plusbytenc b1 b2) = (b1 + b2) \mod 256.
- intros;
- unfold plusbytenc;
- generalize in match (plusbyte_ok b1 b2 false);
- elim (plusbyte b1 b2 false);
- simplify in H ⊢ %;
- change with (nat_of_byte t = (b1 + b2) \mod 256);
- rewrite < plus_n_O in H;
- rewrite > H; clear H;
- rewrite > mod_plus;
- letin K ≝ (eq_nat_of_byte_mod t); clearbody K;
- letin K' ≝ (eq_mod_times_n_m_m_O (nat_of_bool t1) 256 ?); clearbody K';
-  [ autobatch | ];
- autobatch paramodulation.
-qed.
-
-
-
-lemma eq_eqbyte_x0_x0_byte_of_nat_S_false:
- ∀b. b < 255 → eqbyte (mk_byte x0 x0) (byte_of_nat (S b)) = false.
- intros;
- unfold byte_of_nat;
- cut (b < 15 ∨ b ≥ 15);
-  [ elim Hcut;
-    [ unfold eqbyte;
-      change in ⊢ (? ? (? ? %) ?) with (eqex x0 (exadecimal_of_nat (S b))); 
-      rewrite > eq_eqex_S_x0_false;
-       [ elim (eqex (bh (mk_byte x0 x0))
-          (bh (mk_byte (exadecimal_of_nat (S b/16)) (exadecimal_of_nat (S b)))));
-         simplify;
-         reflexivity
-       | assumption
-       ]
-    | unfold eqbyte;
-      change in ⊢ (? ? (? % ?) ?) with (eqex x0 (exadecimal_of_nat (S b/16)));
-      letin K ≝ (leq_m_n_to_eq_div_n_m_S (S b) 16 ? ?);
-       [ autobatch
-       | unfold in H1;
-         apply le_S_S;
-         assumption
-       | clearbody K;
-         elim K; clear K;
-         rewrite > H2;
-         rewrite > eq_eqex_S_x0_false;
-          [ reflexivity
-          | unfold lt;
-            unfold lt in H;
-            rewrite < H2;
-            clear H2; clear a; clear H1; clear Hcut;
-            apply (le_times_to_le 16) [ autobatch | ] ;
-            rewrite > (div_mod (S b) 16) in H;[2:autobatch|]
-            rewrite > (div_mod 255 16) in H:(? ? %);[2:autobatch|]
-            lapply (le_to_le_plus_to_le ? ? ? ? ? H);
-            [apply lt_S_to_le;
-             apply lt_mod_m_m;autobatch
-            |rewrite > sym_times;
-             rewrite > sym_times in ⊢ (? ? %); (* just to speed up qed *)
-             normalize in ⊢ (? ? %);apply Hletin;
-            ]
-          ] 
-       ]
-    ]
-  | elim (or_lt_le b 15);unfold ge;autobatch
-  ].
-qed.
-
-axiom eq_mod_O_to_exists: ∀n,m. n \mod m = 0 → ∃z. n = z*m.
-
-lemma eq_bpred_S_a_a:
- ∀a. a < 255 → bpred (byte_of_nat (S a)) = byte_of_nat a.
- intros;
- unfold bpred;
- apply (bool_elim ? (eqex (bl (byte_of_nat (S a))) x0)); intros;
-  [ change with (mk_byte (xpred (bh (byte_of_nat (S a)))) (xpred (bl (byte_of_nat (S a))))
-     = byte_of_nat a);
-    rewrite > (eqex_true_to_eq ? ? H1);
-    normalize in ⊢ (? ? (? ? %) ?);
-    unfold byte_of_nat;
-    change with (mk_byte (xpred (exadecimal_of_nat (S a/16))) xF =
-                 mk_byte (exadecimal_of_nat (a/16)) (exadecimal_of_nat a));
-    lapply (eqex_true_to_eq ? ? H1); clear H1;
-    unfold byte_of_nat in Hletin;
-    change in Hletin with (exadecimal_of_nat (S a) = x0);
-    lapply (eq_f ? ? nat_of_exadecimal ? ? Hletin); clear Hletin;
-    normalize in Hletin1:(? ? ? %);
-    rewrite > nat_of_exadecimal_exadecimal_of_nat in Hletin1;
-    elim (eq_mod_O_to_exists ? ? Hletin1); clear Hletin1;
-    rewrite > H1;
-    rewrite > lt_O_to_div_times; [2: autobatch | ]
-    lapply (eq_f ? ? (λx.x/16) ? ? H1);
-    rewrite > lt_O_to_div_times in Hletin; [2: autobatch | ]
-    lapply (eq_f ? ? (λx.x \mod 16) ? ? H1);
-    rewrite > eq_mod_times_n_m_m_O in Hletin1;
-    elim daemon
-  | change with (mk_byte (bh (byte_of_nat (S a))) (xpred (bl (byte_of_nat (S a))))
-    = byte_of_nat a);
-    unfold byte_of_nat;
-    change with (mk_byte (exadecimal_of_nat (S a/16)) (xpred (exadecimal_of_nat (S a)))
-    = mk_byte (exadecimal_of_nat (a/16)) (exadecimal_of_nat a));
-    lapply (eqex_false_to_not_eq ? ? H1);
-    unfold byte_of_nat in Hletin;
-    change in Hletin with (exadecimal_of_nat (S a) ≠ x0);
-    cut (nat_of_exadecimal (exadecimal_of_nat (S a)) ≠ 0);
-     [2: intro;
-       apply Hletin;
-       lapply (eq_f ? ? exadecimal_of_nat ? ? H2);
-       rewrite > exadecimal_of_nat_nat_of_exadecimal in Hletin1;
-       apply Hletin1
-     | ];
-     
-    elim daemon
-  ]
-qed.
-
-lemma plusbytenc_S:
- ∀x:byte.∀n.plusbytenc (byte_of_nat (x*n)) x = byte_of_nat (x * S n).
- intros;
- rewrite < byte_of_nat_nat_of_byte;
- rewrite > (plusbytenc_ok (byte_of_nat (x*n)) x);
- rewrite < times_n_Sm;
- rewrite > nat_of_byte_byte_of_nat in ⊢ (? ? (? (? (? % ?) ?)) ?);
- rewrite > eq_nat_of_byte_n_nat_of_byte_mod_n_256 in ⊢ (? ? ? %);
- rewrite > mod_plus in ⊢ (? ? (? %) ?);
- rewrite > mod_plus in ⊢ (? ? ? (? %));
- rewrite < mod_mod in ⊢ (? ? (? (? (? % ?) ?)) ?); [2: autobatch | ];
- rewrite > sym_plus in ⊢ (? ? (? (? % ?)) ?);
- reflexivity.
-qed. 
-
-lemma eq_plusbytec_x0_x0_x_false:
- ∀x.plusbytec (mk_byte x0 x0) x = false.
- intro;
- elim x;
- elim e;
- elim e1;
- reflexivity.
-qed.
diff --git a/helm/software/matita/library/assembly/exadecimal.ma b/helm/software/matita/library/assembly/exadecimal.ma
deleted file mode 100644 (file)
index c059b43..0000000
+++ /dev/null
@@ -1,926 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||M||                                                             *)
-(*      ||A||       A project by Andrea Asperti                           *)
-(*      ||T||                                                             *)
-(*      ||I||       Developers:                                           *)
-(*      ||T||         The HELM team.                                      *)
-(*      ||A||         http://helm.cs.unibo.it                             *)
-(*      \   /                                                             *)
-(*       \ /        This file is distributed under the terms of the       *)
-(*        v         GNU General Public License Version 2                  *)
-(*                                                                        *)
-(**************************************************************************)
-
-set "baseuri" "cic:/matita/assembly/exadecimal/".
-
-include "assembly/extra.ma".
-
-inductive exadecimal : Type ≝
-   x0: exadecimal
- | x1: exadecimal
- | x2: exadecimal
- | x3: exadecimal
- | x4: exadecimal
- | x5: exadecimal
- | x6: exadecimal
- | x7: exadecimal
- | x8: exadecimal
- | x9: exadecimal
- | xA: exadecimal
- | xB: exadecimal
- | xC: exadecimal
- | xD: exadecimal
- | xE: exadecimal
- | xF: exadecimal.
-
-definition eqex ≝
- λb1,b2.
-  match b1 with
-   [ x0 ⇒
-       match b2 with
-        [ x0 ⇒ true  | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
-        | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
-        | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
-        | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
-   | x1 ⇒
-       match b2 with
-        [ x0 ⇒ false | x1 ⇒ true  | x2 ⇒ false | x3 ⇒ false
-        | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
-        | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
-        | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
-   | x2 ⇒
-       match b2 with
-        [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ true  | x3 ⇒ false
-        | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
-        | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
-        | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
-   | x3 ⇒
-       match b2 with
-        [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ true 
-        | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
-        | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
-        | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
-   | x4 ⇒
-       match b2 with
-        [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
-        | x4 ⇒ true  | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
-        | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
-        | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
-   | x5 ⇒
-       match b2 with
-        [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
-        | x4 ⇒ false | x5 ⇒ true  | x6 ⇒ false | x7 ⇒ false
-        | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
-        | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
-   | x6 ⇒
-       match b2 with
-        [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
-        | x4 ⇒ false | x5 ⇒ false | x6 ⇒ true  | x7 ⇒ false
-        | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
-        | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
-   | x7 ⇒
-       match b2 with
-        [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
-        | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ true 
-        | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
-        | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
-   | x8 ⇒
-       match b2 with
-        [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
-        | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
-        | x8 ⇒ true  | x9 ⇒ false | xA ⇒ false | xB ⇒ false
-        | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
-   | x9 ⇒
-       match b2 with
-        [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
-        | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
-        | x8 ⇒ false | x9 ⇒ true  | xA ⇒ false | xB ⇒ false
-        | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
-   | xA ⇒
-       match b2 with
-        [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
-        | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
-        | x8 ⇒ false | x9 ⇒ false | xA ⇒ true  | xB ⇒ false
-        | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
-   | xB ⇒
-       match b2 with
-        [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
-        | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
-        | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ true 
-        | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
-   | xC ⇒
-       match b2 with
-        [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
-        | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
-        | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
-        | xC ⇒ true  | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
-   | xD ⇒
-       match b2 with
-        [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
-        | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
-        | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
-        | xC ⇒ false | xD ⇒ true  | xE ⇒ false | xF ⇒ false ] 
-   | xE ⇒
-       match b2 with
-        [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
-        | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
-        | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
-        | xC ⇒ false | xD ⇒ false | xE ⇒ true  | xF ⇒ false ] 
-   | xF ⇒
-       match b2 with
-        [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
-        | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
-        | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
-        | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ true  ]]. 
-
-definition plusex ≝
- λb1,b2,c.
-  match c with
-   [ true ⇒
-      match b1 with
-       [ x0 ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool x1 false
-            | x1 ⇒ couple exadecimal bool x2 false
-            | x2 ⇒ couple exadecimal bool x3 false
-            | x3 ⇒ couple exadecimal bool x4 false
-            | x4 ⇒ couple exadecimal bool x5 false
-            | x5 ⇒ couple exadecimal bool x6 false
-            | x6 ⇒ couple exadecimal bool x7 false
-            | x7 ⇒ couple exadecimal bool x8 false
-            | x8 ⇒ couple exadecimal bool x9 false
-            | x9 ⇒ couple exadecimal bool xA false
-            | xA ⇒ couple exadecimal bool xB false
-            | xB ⇒ couple exadecimal bool xC false
-            | xC ⇒ couple exadecimal bool xD false
-            | xD ⇒ couple exadecimal bool xE false
-            | xE ⇒ couple exadecimal bool xF false
-            | xF ⇒ couple exadecimal bool x0 true ] 
-       | x1 ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool x2 false
-            | x1 ⇒ couple exadecimal bool x3 false
-            | x2 ⇒ couple exadecimal bool x4 false
-            | x3 ⇒ couple exadecimal bool x5 false
-            | x4 ⇒ couple exadecimal bool x6 false
-            | x5 ⇒ couple exadecimal bool x7 false
-            | x6 ⇒ couple exadecimal bool x8 false
-            | x7 ⇒ couple exadecimal bool x9 false
-            | x8 ⇒ couple exadecimal bool xA false
-            | x9 ⇒ couple exadecimal bool xB false
-            | xA ⇒ couple exadecimal bool xC false
-            | xB ⇒ couple exadecimal bool xD false
-            | xC ⇒ couple exadecimal bool xE false
-            | xD ⇒ couple exadecimal bool xF false
-            | xE ⇒ couple exadecimal bool x0 true
-            | xF ⇒ couple exadecimal bool x1 true ] 
-       | x2 ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool x3 false
-            | x1 ⇒ couple exadecimal bool x4 false
-            | x2 ⇒ couple exadecimal bool x5 false
-            | x3 ⇒ couple exadecimal bool x6 false
-            | x4 ⇒ couple exadecimal bool x7 false
-            | x5 ⇒ couple exadecimal bool x8 false
-            | x6 ⇒ couple exadecimal bool x9 false
-            | x7 ⇒ couple exadecimal bool xA false
-            | x8 ⇒ couple exadecimal bool xB false
-            | x9 ⇒ couple exadecimal bool xC false
-            | xA ⇒ couple exadecimal bool xD false
-            | xB ⇒ couple exadecimal bool xE false
-            | xC ⇒ couple exadecimal bool xF false
-            | xD ⇒ couple exadecimal bool x0 true
-            | xE ⇒ couple exadecimal bool x1 true
-            | xF ⇒ couple exadecimal bool x2 true ] 
-       | x3 ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool x4 false
-            | x1 ⇒ couple exadecimal bool x5 false
-            | x2 ⇒ couple exadecimal bool x6 false
-            | x3 ⇒ couple exadecimal bool x7 false
-            | x4 ⇒ couple exadecimal bool x8 false
-            | x5 ⇒ couple exadecimal bool x9 false
-            | x6 ⇒ couple exadecimal bool xA false
-            | x7 ⇒ couple exadecimal bool xB false
-            | x8 ⇒ couple exadecimal bool xC false
-            | x9 ⇒ couple exadecimal bool xD false
-            | xA ⇒ couple exadecimal bool xE false
-            | xB ⇒ couple exadecimal bool xF false
-            | xC ⇒ couple exadecimal bool x0 true
-            | xD ⇒ couple exadecimal bool x1 true
-            | xE ⇒ couple exadecimal bool x2 true
-            | xF ⇒ couple exadecimal bool x3 true ] 
-       | x4 ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool x5 false
-            | x1 ⇒ couple exadecimal bool x6 false
-            | x2 ⇒ couple exadecimal bool x7 false
-            | x3 ⇒ couple exadecimal bool x8 false
-            | x4 ⇒ couple exadecimal bool x9 false
-            | x5 ⇒ couple exadecimal bool xA false
-            | x6 ⇒ couple exadecimal bool xB false
-            | x7 ⇒ couple exadecimal bool xC false
-            | x8 ⇒ couple exadecimal bool xD false
-            | x9 ⇒ couple exadecimal bool xE false
-            | xA ⇒ couple exadecimal bool xF false
-            | xB ⇒ couple exadecimal bool x0 true
-            | xC ⇒ couple exadecimal bool x1 true
-            | xD ⇒ couple exadecimal bool x2 true
-            | xE ⇒ couple exadecimal bool x3 true
-            | xF ⇒ couple exadecimal bool x4 true ] 
-       | x5 ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool x6 false
-            | x1 ⇒ couple exadecimal bool x7 false
-            | x2 ⇒ couple exadecimal bool x8 false
-            | x3 ⇒ couple exadecimal bool x9 false
-            | x4 ⇒ couple exadecimal bool xA false
-            | x5 ⇒ couple exadecimal bool xB false
-            | x6 ⇒ couple exadecimal bool xC false
-            | x7 ⇒ couple exadecimal bool xD false
-            | x8 ⇒ couple exadecimal bool xE false
-            | x9 ⇒ couple exadecimal bool xF false
-            | xA ⇒ couple exadecimal bool x0 true
-            | xB ⇒ couple exadecimal bool x1 true
-            | xC ⇒ couple exadecimal bool x2 true
-            | xD ⇒ couple exadecimal bool x3 true
-            | xE ⇒ couple exadecimal bool x4 true
-            | xF ⇒ couple exadecimal bool x5 true ] 
-       | x6 ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool x7 false
-            | x1 ⇒ couple exadecimal bool x8 false
-            | x2 ⇒ couple exadecimal bool x9 false
-            | x3 ⇒ couple exadecimal bool xA false
-            | x4 ⇒ couple exadecimal bool xB false
-            | x5 ⇒ couple exadecimal bool xC false
-            | x6 ⇒ couple exadecimal bool xD false
-            | x7 ⇒ couple exadecimal bool xE false
-            | x8 ⇒ couple exadecimal bool xF false
-            | x9 ⇒ couple exadecimal bool x0 true
-            | xA ⇒ couple exadecimal bool x1 true
-            | xB ⇒ couple exadecimal bool x2 true
-            | xC ⇒ couple exadecimal bool x3 true
-            | xD ⇒ couple exadecimal bool x4 true
-            | xE ⇒ couple exadecimal bool x5 true
-            | xF ⇒ couple exadecimal bool x6 true ] 
-       | x7 ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool x8 false
-            | x1 ⇒ couple exadecimal bool x9 false
-            | x2 ⇒ couple exadecimal bool xA false
-            | x3 ⇒ couple exadecimal bool xB false
-            | x4 ⇒ couple exadecimal bool xC false
-            | x5 ⇒ couple exadecimal bool xD false
-            | x6 ⇒ couple exadecimal bool xE false
-            | x7 ⇒ couple exadecimal bool xF false
-            | x8 ⇒ couple exadecimal bool x0 true
-            | x9 ⇒ couple exadecimal bool x1 true
-            | xA ⇒ couple exadecimal bool x2 true
-            | xB ⇒ couple exadecimal bool x3 true
-            | xC ⇒ couple exadecimal bool x4 true
-            | xD ⇒ couple exadecimal bool x5 true
-            | xE ⇒ couple exadecimal bool x6 true
-            | xF ⇒ couple exadecimal bool x7 true ] 
-       | x8 ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool x9 false
-            | x1 ⇒ couple exadecimal bool xA false
-            | x2 ⇒ couple exadecimal bool xB false
-            | x3 ⇒ couple exadecimal bool xC false
-            | x4 ⇒ couple exadecimal bool xD false
-            | x5 ⇒ couple exadecimal bool xE false
-            | x6 ⇒ couple exadecimal bool xF false
-            | x7 ⇒ couple exadecimal bool x0 true
-            | x8 ⇒ couple exadecimal bool x1 true
-            | x9 ⇒ couple exadecimal bool x2 true
-            | xA ⇒ couple exadecimal bool x3 true
-            | xB ⇒ couple exadecimal bool x4 true
-            | xC ⇒ couple exadecimal bool x5 true
-            | xD ⇒ couple exadecimal bool x6 true
-            | xE ⇒ couple exadecimal bool x7 true
-            | xF ⇒ couple exadecimal bool x8 true ] 
-       | x9 ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool xA false
-            | x1 ⇒ couple exadecimal bool xB false
-            | x2 ⇒ couple exadecimal bool xC false
-            | x3 ⇒ couple exadecimal bool xD false
-            | x4 ⇒ couple exadecimal bool xE false
-            | x5 ⇒ couple exadecimal bool xF false
-            | x6 ⇒ couple exadecimal bool x0 true
-            | x7 ⇒ couple exadecimal bool x1 true
-            | x8 ⇒ couple exadecimal bool x2 true
-            | x9 ⇒ couple exadecimal bool x3 true
-            | xA ⇒ couple exadecimal bool x4 true
-            | xB ⇒ couple exadecimal bool x5 true
-            | xC ⇒ couple exadecimal bool x6 true
-            | xD ⇒ couple exadecimal bool x7 true
-            | xE ⇒ couple exadecimal bool x8 true
-            | xF ⇒ couple exadecimal bool x9 true ] 
-       | xA ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool xB false
-            | x1 ⇒ couple exadecimal bool xC false
-            | x2 ⇒ couple exadecimal bool xD false
-            | x3 ⇒ couple exadecimal bool xE false
-            | x4 ⇒ couple exadecimal bool xF false
-            | x5 ⇒ couple exadecimal bool x0 true
-            | x6 ⇒ couple exadecimal bool x1 true
-            | x7 ⇒ couple exadecimal bool x2 true
-            | x8 ⇒ couple exadecimal bool x3 true
-            | x9 ⇒ couple exadecimal bool x4 true
-            | xA ⇒ couple exadecimal bool x5 true
-            | xB ⇒ couple exadecimal bool x6 true
-            | xC ⇒ couple exadecimal bool x7 true
-            | xD ⇒ couple exadecimal bool x8 true
-            | xE ⇒ couple exadecimal bool x9 true
-            | xF ⇒ couple exadecimal bool xA true ] 
-       | xB ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool xC false
-            | x1 ⇒ couple exadecimal bool xD false
-            | x2 ⇒ couple exadecimal bool xE false
-            | x3 ⇒ couple exadecimal bool xF false
-            | x4 ⇒ couple exadecimal bool x0 true
-            | x5 ⇒ couple exadecimal bool x1 true
-            | x6 ⇒ couple exadecimal bool x2 true
-            | x7 ⇒ couple exadecimal bool x3 true
-            | x8 ⇒ couple exadecimal bool x4 true
-            | x9 ⇒ couple exadecimal bool x5 true
-            | xA ⇒ couple exadecimal bool x6 true
-            | xB ⇒ couple exadecimal bool x7 true
-            | xC ⇒ couple exadecimal bool x8 true
-            | xD ⇒ couple exadecimal bool x9 true
-            | xE ⇒ couple exadecimal bool xA true
-            | xF ⇒ couple exadecimal bool xB true ] 
-       | xC ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool xD false
-            | x1 ⇒ couple exadecimal bool xE false
-            | x2 ⇒ couple exadecimal bool xF false
-            | x3 ⇒ couple exadecimal bool x0 true
-            | x4 ⇒ couple exadecimal bool x1 true
-            | x5 ⇒ couple exadecimal bool x2 true
-            | x6 ⇒ couple exadecimal bool x3 true
-            | x7 ⇒ couple exadecimal bool x4 true
-            | x8 ⇒ couple exadecimal bool x5 true
-            | x9 ⇒ couple exadecimal bool x6 true
-            | xA ⇒ couple exadecimal bool x7 true
-            | xB ⇒ couple exadecimal bool x8 true
-            | xC ⇒ couple exadecimal bool x9 true
-            | xD ⇒ couple exadecimal bool xA true
-            | xE ⇒ couple exadecimal bool xB true
-            | xF ⇒ couple exadecimal bool xC true ] 
-       | xD ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool xE false
-            | x1 ⇒ couple exadecimal bool xF false
-            | x2 ⇒ couple exadecimal bool x0 true
-            | x3 ⇒ couple exadecimal bool x1 true
-            | x4 ⇒ couple exadecimal bool x2 true
-            | x5 ⇒ couple exadecimal bool x3 true
-            | x6 ⇒ couple exadecimal bool x4 true
-            | x7 ⇒ couple exadecimal bool x5 true
-            | x8 ⇒ couple exadecimal bool x6 true
-            | x9 ⇒ couple exadecimal bool x7 true
-            | xA ⇒ couple exadecimal bool x8 true
-            | xB ⇒ couple exadecimal bool x9 true
-            | xC ⇒ couple exadecimal bool xA true
-            | xD ⇒ couple exadecimal bool xB true
-            | xE ⇒ couple exadecimal bool xC true
-            | xF ⇒ couple exadecimal bool xD true ] 
-       | xE ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool xF false
-            | x1 ⇒ couple exadecimal bool x0 true
-            | x2 ⇒ couple exadecimal bool x1 true
-            | x3 ⇒ couple exadecimal bool x2 true
-            | x4 ⇒ couple exadecimal bool x3 true
-            | x5 ⇒ couple exadecimal bool x4 true
-            | x6 ⇒ couple exadecimal bool x5 true
-            | x7 ⇒ couple exadecimal bool x6 true
-            | x8 ⇒ couple exadecimal bool x7 true
-            | x9 ⇒ couple exadecimal bool x8 true
-            | xA ⇒ couple exadecimal bool x9 true
-            | xB ⇒ couple exadecimal bool xA true
-            | xC ⇒ couple exadecimal bool xB true
-            | xD ⇒ couple exadecimal bool xC true
-            | xE ⇒ couple exadecimal bool xD true
-            | xF ⇒ couple exadecimal bool xE true ]
-       | xF ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool x0 true
-            | x1 ⇒ couple exadecimal bool x1 true
-            | x2 ⇒ couple exadecimal bool x2 true
-            | x3 ⇒ couple exadecimal bool x3 true
-            | x4 ⇒ couple exadecimal bool x4 true
-            | x5 ⇒ couple exadecimal bool x5 true
-            | x6 ⇒ couple exadecimal bool x6 true
-            | x7 ⇒ couple exadecimal bool x7 true
-            | x8 ⇒ couple exadecimal bool x8 true
-            | x9 ⇒ couple exadecimal bool x9 true
-            | xA ⇒ couple exadecimal bool xA true
-            | xB ⇒ couple exadecimal bool xB true
-            | xC ⇒ couple exadecimal bool xC true
-            | xD ⇒ couple exadecimal bool xD true
-            | xE ⇒ couple exadecimal bool xE true
-            | xF ⇒ couple exadecimal bool xF true ] 
-       ]
-   | false ⇒
-      match b1 with
-       [ x0 ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool x0 false
-            | x1 ⇒ couple exadecimal bool x1 false
-            | x2 ⇒ couple exadecimal bool x2 false
-            | x3 ⇒ couple exadecimal bool x3 false
-            | x4 ⇒ couple exadecimal bool x4 false
-            | x5 ⇒ couple exadecimal bool x5 false
-            | x6 ⇒ couple exadecimal bool x6 false
-            | x7 ⇒ couple exadecimal bool x7 false
-            | x8 ⇒ couple exadecimal bool x8 false
-            | x9 ⇒ couple exadecimal bool x9 false
-            | xA ⇒ couple exadecimal bool xA false
-            | xB ⇒ couple exadecimal bool xB false
-            | xC ⇒ couple exadecimal bool xC false
-            | xD ⇒ couple exadecimal bool xD false
-            | xE ⇒ couple exadecimal bool xE false
-            | xF ⇒ couple exadecimal bool xF false ] 
-       | x1 ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool x1 false
-            | x1 ⇒ couple exadecimal bool x2 false
-            | x2 ⇒ couple exadecimal bool x3 false
-            | x3 ⇒ couple exadecimal bool x4 false
-            | x4 ⇒ couple exadecimal bool x5 false
-            | x5 ⇒ couple exadecimal bool x6 false
-            | x6 ⇒ couple exadecimal bool x7 false
-            | x7 ⇒ couple exadecimal bool x8 false
-            | x8 ⇒ couple exadecimal bool x9 false
-            | x9 ⇒ couple exadecimal bool xA false
-            | xA ⇒ couple exadecimal bool xB false
-            | xB ⇒ couple exadecimal bool xC false
-            | xC ⇒ couple exadecimal bool xD false
-            | xD ⇒ couple exadecimal bool xE false
-            | xE ⇒ couple exadecimal bool xF false
-            | xF ⇒ couple exadecimal bool x0 true ] 
-       | x2 ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool x2 false
-            | x1 ⇒ couple exadecimal bool x3 false
-            | x2 ⇒ couple exadecimal bool x4 false
-            | x3 ⇒ couple exadecimal bool x5 false
-            | x4 ⇒ couple exadecimal bool x6 false
-            | x5 ⇒ couple exadecimal bool x7 false
-            | x6 ⇒ couple exadecimal bool x8 false
-            | x7 ⇒ couple exadecimal bool x9 false
-            | x8 ⇒ couple exadecimal bool xA false
-            | x9 ⇒ couple exadecimal bool xB false
-            | xA ⇒ couple exadecimal bool xC false
-            | xB ⇒ couple exadecimal bool xD false
-            | xC ⇒ couple exadecimal bool xE false
-            | xD ⇒ couple exadecimal bool xF false
-            | xE ⇒ couple exadecimal bool x0 true
-            | xF ⇒ couple exadecimal bool x1 true ] 
-       | x3 ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool x3 false
-            | x1 ⇒ couple exadecimal bool x4 false
-            | x2 ⇒ couple exadecimal bool x5 false
-            | x3 ⇒ couple exadecimal bool x6 false
-            | x4 ⇒ couple exadecimal bool x7 false
-            | x5 ⇒ couple exadecimal bool x8 false
-            | x6 ⇒ couple exadecimal bool x9 false
-            | x7 ⇒ couple exadecimal bool xA false
-            | x8 ⇒ couple exadecimal bool xB false
-            | x9 ⇒ couple exadecimal bool xC false
-            | xA ⇒ couple exadecimal bool xD false
-            | xB ⇒ couple exadecimal bool xE false
-            | xC ⇒ couple exadecimal bool xF false
-            | xD ⇒ couple exadecimal bool x0 true
-            | xE ⇒ couple exadecimal bool x1 true
-            | xF ⇒ couple exadecimal bool x2 true ] 
-       | x4 ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool x4 false
-            | x1 ⇒ couple exadecimal bool x5 false
-            | x2 ⇒ couple exadecimal bool x6 false
-            | x3 ⇒ couple exadecimal bool x7 false
-            | x4 ⇒ couple exadecimal bool x8 false
-            | x5 ⇒ couple exadecimal bool x9 false
-            | x6 ⇒ couple exadecimal bool xA false
-            | x7 ⇒ couple exadecimal bool xB false
-            | x8 ⇒ couple exadecimal bool xC false
-            | x9 ⇒ couple exadecimal bool xD false
-            | xA ⇒ couple exadecimal bool xE false
-            | xB ⇒ couple exadecimal bool xF false
-            | xC ⇒ couple exadecimal bool x0 true
-            | xD ⇒ couple exadecimal bool x1 true
-            | xE ⇒ couple exadecimal bool x2 true
-            | xF ⇒ couple exadecimal bool x3 true ] 
-       | x5 ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool x5 false
-            | x1 ⇒ couple exadecimal bool x6 false
-            | x2 ⇒ couple exadecimal bool x7 false
-            | x3 ⇒ couple exadecimal bool x8 false
-            | x4 ⇒ couple exadecimal bool x9 false
-            | x5 ⇒ couple exadecimal bool xA false
-            | x6 ⇒ couple exadecimal bool xB false
-            | x7 ⇒ couple exadecimal bool xC false
-            | x8 ⇒ couple exadecimal bool xD false
-            | x9 ⇒ couple exadecimal bool xE false
-            | xA ⇒ couple exadecimal bool xF false
-            | xB ⇒ couple exadecimal bool x0 true
-            | xC ⇒ couple exadecimal bool x1 true
-            | xD ⇒ couple exadecimal bool x2 true
-            | xE ⇒ couple exadecimal bool x3 true
-            | xF ⇒ couple exadecimal bool x4 true ] 
-       | x6 ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool x6 false
-            | x1 ⇒ couple exadecimal bool x7 false
-            | x2 ⇒ couple exadecimal bool x8 false
-            | x3 ⇒ couple exadecimal bool x9 false
-            | x4 ⇒ couple exadecimal bool xA false
-            | x5 ⇒ couple exadecimal bool xB false
-            | x6 ⇒ couple exadecimal bool xC false
-            | x7 ⇒ couple exadecimal bool xD false
-            | x8 ⇒ couple exadecimal bool xE false
-            | x9 ⇒ couple exadecimal bool xF false
-            | xA ⇒ couple exadecimal bool x0 true
-            | xB ⇒ couple exadecimal bool x1 true
-            | xC ⇒ couple exadecimal bool x2 true
-            | xD ⇒ couple exadecimal bool x3 true
-            | xE ⇒ couple exadecimal bool x4 true
-            | xF ⇒ couple exadecimal bool x5 true ] 
-       | x7 ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool x7 false
-            | x1 ⇒ couple exadecimal bool x8 false
-            | x2 ⇒ couple exadecimal bool x9 false
-            | x3 ⇒ couple exadecimal bool xA false
-            | x4 ⇒ couple exadecimal bool xB false
-            | x5 ⇒ couple exadecimal bool xC false
-            | x6 ⇒ couple exadecimal bool xD false
-            | x7 ⇒ couple exadecimal bool xE false
-            | x8 ⇒ couple exadecimal bool xF false
-            | x9 ⇒ couple exadecimal bool x0 true
-            | xA ⇒ couple exadecimal bool x1 true
-            | xB ⇒ couple exadecimal bool x2 true
-            | xC ⇒ couple exadecimal bool x3 true
-            | xD ⇒ couple exadecimal bool x4 true
-            | xE ⇒ couple exadecimal bool x5 true
-            | xF ⇒ couple exadecimal bool x6 true ] 
-       | x8 ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool x8 false
-            | x1 ⇒ couple exadecimal bool x9 false
-            | x2 ⇒ couple exadecimal bool xA false
-            | x3 ⇒ couple exadecimal bool xB false
-            | x4 ⇒ couple exadecimal bool xC false
-            | x5 ⇒ couple exadecimal bool xD false
-            | x6 ⇒ couple exadecimal bool xE false
-            | x7 ⇒ couple exadecimal bool xF false
-            | x8 ⇒ couple exadecimal bool x0 true
-            | x9 ⇒ couple exadecimal bool x1 true
-            | xA ⇒ couple exadecimal bool x2 true
-            | xB ⇒ couple exadecimal bool x3 true
-            | xC ⇒ couple exadecimal bool x4 true
-            | xD ⇒ couple exadecimal bool x5 true
-            | xE ⇒ couple exadecimal bool x6 true
-            | xF ⇒ couple exadecimal bool x7 true ] 
-       | x9 ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool x9 false
-            | x1 ⇒ couple exadecimal bool xA false
-            | x2 ⇒ couple exadecimal bool xB false
-            | x3 ⇒ couple exadecimal bool xC false
-            | x4 ⇒ couple exadecimal bool xD false
-            | x5 ⇒ couple exadecimal bool xE false
-            | x6 ⇒ couple exadecimal bool xF false
-            | x7 ⇒ couple exadecimal bool x0 true
-            | x8 ⇒ couple exadecimal bool x1 true
-            | x9 ⇒ couple exadecimal bool x2 true
-            | xA ⇒ couple exadecimal bool x3 true
-            | xB ⇒ couple exadecimal bool x4 true
-            | xC ⇒ couple exadecimal bool x5 true
-            | xD ⇒ couple exadecimal bool x6 true
-            | xE ⇒ couple exadecimal bool x7 true
-            | xF ⇒ couple exadecimal bool x8 true ] 
-       | xA ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool xA false
-            | x1 ⇒ couple exadecimal bool xB false
-            | x2 ⇒ couple exadecimal bool xC false
-            | x3 ⇒ couple exadecimal bool xD false
-            | x4 ⇒ couple exadecimal bool xE false
-            | x5 ⇒ couple exadecimal bool xF false
-            | x6 ⇒ couple exadecimal bool x0 true
-            | x7 ⇒ couple exadecimal bool x1 true
-            | x8 ⇒ couple exadecimal bool x2 true
-            | x9 ⇒ couple exadecimal bool x3 true
-            | xA ⇒ couple exadecimal bool x4 true
-            | xB ⇒ couple exadecimal bool x5 true
-            | xC ⇒ couple exadecimal bool x6 true
-            | xD ⇒ couple exadecimal bool x7 true
-            | xE ⇒ couple exadecimal bool x8 true
-            | xF ⇒ couple exadecimal bool x9 true ] 
-       | xB ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool xB false
-            | x1 ⇒ couple exadecimal bool xC false
-            | x2 ⇒ couple exadecimal bool xD false
-            | x3 ⇒ couple exadecimal bool xE false
-            | x4 ⇒ couple exadecimal bool xF false
-            | x5 ⇒ couple exadecimal bool x0 true
-            | x6 ⇒ couple exadecimal bool x1 true
-            | x7 ⇒ couple exadecimal bool x2 true
-            | x8 ⇒ couple exadecimal bool x3 true
-            | x9 ⇒ couple exadecimal bool x4 true
-            | xA ⇒ couple exadecimal bool x5 true
-            | xB ⇒ couple exadecimal bool x6 true
-            | xC ⇒ couple exadecimal bool x7 true
-            | xD ⇒ couple exadecimal bool x8 true
-            | xE ⇒ couple exadecimal bool x9 true
-            | xF ⇒ couple exadecimal bool xA true ] 
-       | xC ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool xC false
-            | x1 ⇒ couple exadecimal bool xD false
-            | x2 ⇒ couple exadecimal bool xE false
-            | x3 ⇒ couple exadecimal bool xF false
-            | x4 ⇒ couple exadecimal bool x0 true
-            | x5 ⇒ couple exadecimal bool x1 true
-            | x6 ⇒ couple exadecimal bool x2 true
-            | x7 ⇒ couple exadecimal bool x3 true
-            | x8 ⇒ couple exadecimal bool x4 true
-            | x9 ⇒ couple exadecimal bool x5 true
-            | xA ⇒ couple exadecimal bool x6 true
-            | xB ⇒ couple exadecimal bool x7 true
-            | xC ⇒ couple exadecimal bool x8 true
-            | xD ⇒ couple exadecimal bool x9 true
-            | xE ⇒ couple exadecimal bool xA true
-            | xF ⇒ couple exadecimal bool xB true ] 
-       | xD ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool xD false
-            | x1 ⇒ couple exadecimal bool xE false
-            | x2 ⇒ couple exadecimal bool xF false
-            | x3 ⇒ couple exadecimal bool x0 true
-            | x4 ⇒ couple exadecimal bool x1 true
-            | x5 ⇒ couple exadecimal bool x2 true
-            | x6 ⇒ couple exadecimal bool x3 true
-            | x7 ⇒ couple exadecimal bool x4 true
-            | x8 ⇒ couple exadecimal bool x5 true
-            | x9 ⇒ couple exadecimal bool x6 true
-            | xA ⇒ couple exadecimal bool x7 true
-            | xB ⇒ couple exadecimal bool x8 true
-            | xC ⇒ couple exadecimal bool x9 true
-            | xD ⇒ couple exadecimal bool xA true
-            | xE ⇒ couple exadecimal bool xB true
-            | xF ⇒ couple exadecimal bool xC true ] 
-       | xE ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool xE false
-            | x1 ⇒ couple exadecimal bool xF false
-            | x2 ⇒ couple exadecimal bool x0 true
-            | x3 ⇒ couple exadecimal bool x1 true
-            | x4 ⇒ couple exadecimal bool x2 true
-            | x5 ⇒ couple exadecimal bool x3 true
-            | x6 ⇒ couple exadecimal bool x4 true
-            | x7 ⇒ couple exadecimal bool x5 true
-            | x8 ⇒ couple exadecimal bool x6 true
-            | x9 ⇒ couple exadecimal bool x7 true
-            | xA ⇒ couple exadecimal bool x8 true
-            | xB ⇒ couple exadecimal bool x9 true
-            | xC ⇒ couple exadecimal bool xA true
-            | xD ⇒ couple exadecimal bool xB true
-            | xE ⇒ couple exadecimal bool xC true
-            | xF ⇒ couple exadecimal bool xD true ] 
-       | xF ⇒
-           match b2 with
-            [ x0 ⇒ couple exadecimal bool xF false
-            | x1 ⇒ couple exadecimal bool x0 true
-            | x2 ⇒ couple exadecimal bool x1 true
-            | x3 ⇒ couple exadecimal bool x2 true
-            | x4 ⇒ couple exadecimal bool x3 true
-            | x5 ⇒ couple exadecimal bool x4 true
-            | x6 ⇒ couple exadecimal bool x5 true
-            | x7 ⇒ couple exadecimal bool x6 true
-            | x8 ⇒ couple exadecimal bool x7 true
-            | x9 ⇒ couple exadecimal bool x8 true
-            | xA ⇒ couple exadecimal bool x9 true
-            | xB ⇒ couple exadecimal bool xA true
-            | xC ⇒ couple exadecimal bool xB true
-            | xD ⇒ couple exadecimal bool xC true
-            | xE ⇒ couple exadecimal bool xD true
-            | xF ⇒ couple exadecimal bool xE true ]
-       ]
-   ]
-.
-
-definition nat_of_exadecimal ≝
- λb.
-  match b with
-   [ x0 ⇒ 0
-   | x1 ⇒ 1
-   | x2 ⇒ 2
-   | x3 ⇒ 3
-   | x4 ⇒ 4
-   | x5 ⇒ 5
-   | x6 ⇒ 6
-   | x7 ⇒ 7
-   | x8 ⇒ 8
-   | x9 ⇒ 9
-   | xA ⇒ 10
-   | xB ⇒ 11
-   | xC ⇒ 12
-   | xD ⇒ 13
-   | xE ⇒ 14
-   | xF ⇒ 15
-   ].
-
-coercion cic:/matita/assembly/exadecimal/nat_of_exadecimal.con.
-
-let rec exadecimal_of_nat b ≝
-  match b with [ O ⇒ x0 | S b ⇒
-  match b with [ O ⇒ x1 | S b ⇒
-  match b with [ O ⇒ x2 | S b ⇒ 
-  match b with [ O ⇒ x3 | S b ⇒ 
-  match b with [ O ⇒ x4 | S b ⇒ 
-  match b with [ O ⇒ x5 | S b ⇒ 
-  match b with [ O ⇒ x6 | S b ⇒ 
-  match b with [ O ⇒ x7 | S b ⇒ 
-  match b with [ O ⇒ x8 | S b ⇒ 
-  match b with [ O ⇒ x9 | S b ⇒ 
-  match b with [ O ⇒ xA | S b ⇒ 
-  match b with [ O ⇒ xB | S b ⇒ 
-  match b with [ O ⇒ xC | S b ⇒ 
-  match b with [ O ⇒ xD | S b ⇒ 
-  match b with [ O ⇒ xE | S b ⇒ 
-  match b with [ O ⇒ xF | S b ⇒ exadecimal_of_nat b ]]]]]]]]]]]]]]]]. 
-
-lemma lt_nat_of_exadecimal_16: ∀b. nat_of_exadecimal b < 16.
- intro;
- elim b;
- simplify;
- [
- |*: repeat (apply lt_to_lt_S_S)
- ];
- autobatch.
-qed.
-
-lemma exadecimal_of_nat_mod:
- ∀n.exadecimal_of_nat n = exadecimal_of_nat (n \mod 16).
- intros;
- apply (nat_elim1 n); intro;
- cases m; [ intro; reflexivity | ];
- cases n1; [ intro; reflexivity | ];
- cases n2; [ intro; reflexivity | ];
- cases n3; [ intro; reflexivity | ];
- cases n4; [ intro; reflexivity | ];
- cases n5; [ intro; reflexivity | ];
- cases n6; [ intro; reflexivity | ];
- cases n7; [ intro; reflexivity | ];
- cases n8; [ intro; reflexivity | ];
- cases n9; [ intro; reflexivity | ];
- cases n10; [ intro; reflexivity | ];
- cases n11; [ intro; reflexivity | ];
- cases n12; [ intro; reflexivity | ];
- cases n13; [ intro; reflexivity | ];
- cases n14; [ intro; reflexivity | ];
- cases n15; [ intro; reflexivity | ];
- intros;
- change in ⊢ (? ? % ?) with (exadecimal_of_nat n16);
- change in ⊢ (? ? ? (? (? % ?))) with (16 + n16);
- rewrite > mod_plus;
- change in ⊢ (? ? ? (? (? % ?))) with (n16 \mod 16);
- rewrite < mod_mod;
-  [ apply H;
-    unfold lt;
-    autobatch.
-  | autobatch
-  ]
-qed.
-
-lemma nat_of_exadecimal_exadecimal_of_nat:
- ∀n. nat_of_exadecimal (exadecimal_of_nat n) = n \mod 16.
- intro;
- rewrite > exadecimal_of_nat_mod;
- generalize in match (lt_mod_m_m n 16 ?); [2: autobatch | ]
- generalize in match (n \mod 16); intro;
- cases n1; [ intro; reflexivity | ];
- cases n2; [ intro; reflexivity | ];
- cases n3; [ intro; reflexivity | ];
- cases n4; [ intro; reflexivity | ];
- cases n5; [ intro; reflexivity | ]; 
- cases n6; [ intro; reflexivity | ]; 
- cases n7; [ intro; reflexivity | ];
- cases n8; [ intro; reflexivity | ];
- cases n9; [ intro; reflexivity | ];
- cases n10; [ intro; reflexivity | ];
- cases n11 [ intro; reflexivity | ];
- cases n12; [ intro; reflexivity | ];
- cases n13; [ intro; reflexivity | ];
- cases n14; [ intro; reflexivity | ];
- cases n15; [ intro; reflexivity | ];
- cases n16; [ intro; reflexivity | ];
- intro;
- unfold lt in H;
- cut False;
-  [ elim Hcut
-  | autobatch
-  ]
-qed.
-
-lemma exadecimal_of_nat_nat_of_exadecimal:
- ∀b.exadecimal_of_nat (nat_of_exadecimal b) = b.
- intro;
- elim b;
- reflexivity.
-qed.
-
-lemma plusex_ok:
- ∀b1,b2,c.
-  match plusex b1 b2 c with
-   [ couple r c' ⇒ b1 + b2 + nat_of_bool c = nat_of_exadecimal r + nat_of_bool c' * 16 ].
- intros;
- elim b1; (elim b2; (elim c; normalize; reflexivity)).
-qed.
-
-definition xpred ≝
- λb.
-  match b with
-   [ x0 ⇒ xF
-   | x1 ⇒ x0
-   | x2 ⇒ x1
-   | x3 ⇒ x2
-   | x4 ⇒ x3
-   | x5 ⇒ x4
-   | x6 ⇒ x5
-   | x7 ⇒ x6
-   | x8 ⇒ x7
-   | x9 ⇒ x8
-   | xA ⇒ x9
-   | xB ⇒ xA
-   | xC ⇒ xB
-   | xD ⇒ xC
-   | xE ⇒ xD
-   | xF ⇒ xE ].
-
-lemma eq_eqex_S_x0_false:
- ∀n. n < 15 → eqex x0 (exadecimal_of_nat (S n)) = false.
- intro;
- cases n 0; [ intro; simplify; reflexivity | clear n];
- cases n1 0; [ intro; simplify; reflexivity | clear n1];
- cases n 0; [ intro; simplify; reflexivity | clear n];
- cases n1 0; [ intro; simplify; reflexivity | clear n1];
- cases n 0; [ intro; simplify; reflexivity | clear n];
- cases n1 0; [ intro; simplify; reflexivity | clear n1];
- cases n 0; [ intro; simplify; reflexivity | clear n];
- cases n1 0; [ intro; simplify; reflexivity | clear n1];
- cases n 0; [ intro; simplify; reflexivity | clear n];
- cases n1 0; [ intro; simplify; reflexivity | clear n1];
- cases n 0; [ intro; simplify; reflexivity | clear n];
- cases n1 0; [ intro; simplify; reflexivity | clear n1];
- cases n 0; [ intro; simplify; reflexivity | clear n];
- cases n1 0; [ intro; simplify; reflexivity | clear n1];
- cases n 0; [ intro; simplify; reflexivity | clear n];
- intro;
- unfold lt in H;
- cut (S n1 ≤ 0);
-  [ elim (not_le_Sn_O ? Hcut)
-  | do 15 (apply le_S_S_to_le);
-    assumption
-  ]
-qed.
-
-lemma eqex_true_to_eq: ∀b,b'. eqex b b' = true → b=b'.
- intros 2;
- elim b 0;
- elim b' 0;
- simplify;
- intro;
- first [ reflexivity | destruct H ].
-qed.
-lemma eqex_false_to_not_eq: ∀b,b'. eqex b b' = false → b ≠ b'.
- intros 2;
- elim b 0;
- elim b' 0;
- simplify;
- intro;
- try (destruct H);
- intro K;
- destruct K.
-qed.
-(*
-lemma xpred_S: ∀b. xpred (exadecimal_of_nat (S b)) = exadecimal_of_nat b.
- intros;
- rewrite > exadecimal_of_nat_mod;
- rewrite > exadecimal_of_nat_mod in ⊢ (? ? ? %);
-*)
diff --git a/helm/software/matita/library/assembly/extra.ma b/helm/software/matita/library/assembly/extra.ma
deleted file mode 100644 (file)
index be15c75..0000000
+++ /dev/null
@@ -1,84 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||M||                                                             *)
-(*      ||A||       A project by Andrea Asperti                           *)
-(*      ||T||                                                             *)
-(*      ||I||       Developers:                                           *)
-(*      ||T||         The HELM team.                                      *)
-(*      ||A||         http://helm.cs.unibo.it                             *)
-(*      \   /                                                             *)
-(*       \ /        This file is distributed under the terms of the       *)
-(*        v         GNU General Public License Version 2                  *)
-(*                                                                        *)
-(**************************************************************************)
-
-set "baseuri" "cic:/matita/assembly/extra".
-
-include "nat/div_and_mod.ma".
-include "nat/primes.ma".
-include "list/list.ma".
-
-axiom mod_plus: ∀a,b,m. (a + b) \mod m = (a \mod m + b \mod m) \mod m.
-axiom mod_mod: ∀a,n,m. n∣m → a \mod n = a \mod n \mod m.
-axiom eq_mod_times_n_m_m_O: ∀n,m. O < m → n * m \mod m = O.
-axiom eq_mod_to_eq_plus_mod: ∀a,b,c,m. a \mod m = b \mod m → (a+c) \mod m = (b+c) \mod m.
-axiom eq_mod_times_times_mod: ∀a,b,n,m. m = a*n → (a*b) \mod m = a * (b \mod n).
-axiom divides_to_eq_mod_mod_mod: ∀a,n,m. n∣m → a \mod m \mod n = a \mod n.
-axiom le_to_le_plus_to_le : ∀a,b,c,d.b\leq d\rarr a+b\leq c+d\rarr a\leq c.
-axiom or_lt_le : ∀n,m. n < m ∨ m ≤ n.
-
-inductive cartesian_product (A,B: Type) : Type ≝
- couple: ∀a:A.∀b:B. cartesian_product A B.
-
-lemma le_to_lt: ∀n,m. n ≤ m → n < S m.
- intros;
- autobatch.
-qed.
-
-alias num (instance 0) = "natural number".
-definition nat_of_bool ≝
- λb. match b with [ true ⇒ 1 | false ⇒ 0 ].
-
-theorem lt_trans: ∀x,y,z. x < y → y < z → x < z.
- unfold lt;
- intros;
- autobatch.
-qed.
-
-lemma leq_m_n_to_eq_div_n_m_S: ∀n,m:nat. 0 < m → m ≤ n → ∃z. n/m = S z.
- intros;
- unfold div;
- apply (ex_intro ? ? (div_aux (pred n) (n-m) (pred m)));
- cut (∃w.m = S w);
-  [ elim Hcut;
-    rewrite > H2;
-    rewrite > H2 in H1;
-    clear Hcut; clear H2; clear H; (*clear m;*)
-    simplify;
-    unfold in ⊢ (? ? % ?);
-    cut (∃z.n = S z);
-     [ elim Hcut; clear Hcut;
-       rewrite > H in H1;
-       rewrite > H; clear m;
-       change in ⊢ (? ? % ?)  with
-        (match leb (S a1) a with
-         [ true ⇒ O
-         | false ⇒ S (div_aux a1 ((S a1) - S a) a)]);
-       cut (S a1 ≰ a);
-        [ apply (leb_elim (S a1) a);
-           [ intro;
-             elim (Hcut H2)
-           | intro;
-             simplify;
-             reflexivity
-           ]
-        | intro;
-          autobatch
-        ]
-     | elim H1; autobatch
-     ]
-  | autobatch
-  ].
-qed.
-
-axiom daemon: False.
diff --git a/helm/software/matita/library/assembly/test.ma b/helm/software/matita/library/assembly/test.ma
deleted file mode 100644 (file)
index 644ceba..0000000
+++ /dev/null
@@ -1,289 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||M||                                                             *)
-(*      ||A||       A project by Andrea Asperti                           *)
-(*      ||T||                                                             *)
-(*      ||I||       Developers:                                           *)
-(*      ||T||         The HELM team.                                      *)
-(*      ||A||         http://helm.cs.unibo.it                             *)
-(*      \   /                                                             *)
-(*       \ /        This file is distributed under the terms of the       *)
-(*        v         GNU General Public License Version 2                  *)
-(*                                                                        *)
-(**************************************************************************)
-
-set "baseuri" "cic:/matita/assembly/test/".
-
-include "assembly/vm.ma".
-
-definition mult_source : list byte ≝
-  [#LDAi; 〈x0, x0〉; (* A := 0 *)
-   #STAd; 〈x2, x0〉; (* Z := A *)
-   #LDAd; 〈x1, xF〉; (* (l1) A := Y *)
-   #BEQ;  〈x0, xA〉; (* if A == 0 then goto l2 *)
-   #LDAd; 〈x2, x0〉; (* A := Z *)
-   #DECd; 〈x1, xF〉; (* Y := Y - 1 *)
-   #ADDd; 〈x1, xE〉; (* A += X *)
-   #STAd; 〈x2, x0〉; (* Z := A *)
-   #BRA;  〈xF, x2〉; (* goto l1 *)
-   #LDAd; 〈x2, x0〉].(* (l2) *)
-
-definition mult_memory ≝
- λx,y.λa:addr.
-     match leb a 29 with
-      [ true ⇒ nth ? mult_source 〈x0, x0〉 a
-      | false ⇒
-         match eqb a 30 with
-          [ true ⇒ x
-          | false ⇒ y
-          ]
-      ].
-
-definition mult_status ≝
- λx,y.
-  mk_status 〈x0, x0〉 0 0 false false (mult_memory x y) 0.
-
-notation " 'M' \sub (x y)" non associative with precedence 80 for 
- @{ 'memory $x $y }.
-interpretation "mult_memory" 'memory x y = 
-  (cic:/matita/assembly/test/mult_memory.con x y).
-
-notation " 'M' \sub (x y) \nbsp a" non associative with precedence 80 for 
- @{ 'memory4 $x $y $a }.
-interpretation "mult_memory4" 'memory4 x y a = 
-  (cic:/matita/assembly/test/mult_memory.con x y a).
-  
-notation " \Sigma \sub (x y)" non associative with precedence 80 for 
- @{ 'status $x $y }.
-
-interpretation "mult_status" 'status x y =
-  (cic:/matita/assembly/test/mult_status.con x y).
-
-lemma test_O_O:
-  let i ≝ 14 in
-  let s ≝ execute (mult_status 〈x0, x0〉 〈x0, x0〉) i in
-   pc s = 20 ∧ mem s 32 = byte_of_nat 0.
- split;
- reflexivity.
-qed.
-
-lemma test_0_2:
-  let x ≝ 〈x0, x0〉 in
-  let y ≝ 〈x0, x2〉 in
-  let i ≝ 14 + 23 * nat_of_byte y in
-  let s ≝ execute (mult_status x y) i in
-   pc s = 20 ∧ mem s 32 = plusbytenc x x.
- intros;
- split;
- reflexivity.
-qed.
-
-lemma test_x_1:
- ∀x.
-  let y ≝ 〈x0, x1〉 in
-  let i ≝ 14 + 23 * nat_of_byte y in
-  let s ≝ execute (mult_status x y) i in
-   pc s = 20 ∧ mem s 32 = x.
- intros;
- split;
-  [ reflexivity
-  | change in ⊢ (? ? % ?) with (plusbytenc 〈x0, x0〉 x);
-    rewrite > plusbytenc_O_x;
-    reflexivity
-  ].
-qed.
-
-lemma test_x_2:
- ∀x.
-  let y ≝ 〈x0, x2〉 in
-  let i ≝ 14 + 23 * nat_of_byte y in
-  let s ≝ execute (mult_status x y) i in
-   pc s = 20 ∧ mem s 32 = plusbytenc x x.
- intros;
- split;
-  [ reflexivity
-  | change in ⊢ (? ? % ?) with
-     (plusbytenc (plusbytenc 〈x0, x0〉 x) x);
-    rewrite > plusbytenc_O_x;
-    reflexivity
-  ].
-qed.
-
-lemma loop_invariant':
- ∀x,y:byte.∀j:nat. j ≤ y →
-  execute (mult_status x y) (5 + 23*j)
-   =
-    mk_status (byte_of_nat (x * j)) 4 0 (eqbyte 〈x0, x0〉 (byte_of_nat (x*j)))
-     (plusbytec (byte_of_nat (x*pred j)) x)
-     (update (update (update (mult_memory x y) 30 x) 31 (byte_of_nat (y - j))) 32
-      (byte_of_nat (x * j)))
-     0.
- intros 3;
- elim j;
-  [ do 2 (rewrite < times_n_O);
-    apply status_eq;
-    [1,2,3,4,7: normalize; reflexivity
-    | rewrite > eq_plusbytec_x0_x0_x_false;
-      normalize;
-      reflexivity 
-    | intro;
-      rewrite < minus_n_O;
-      normalize in ⊢ (? ? (? (? ? %) ?) ?);
-      change in ⊢ (? ? % ?) with (update (mult_memory x y) 32 〈x0, x0〉 a);
-      simplify in ⊢ (? ? ? %);
-      change in ⊢ (? ? ? (? (? (? ? ? %) ? ?) ? ? ?)) with (mult_memory x y 30);
-      rewrite > byte_of_nat_nat_of_byte;
-      change in ⊢ (? ? ? (? (? ? ? %) ? ? ?)) with (mult_memory x y 31);
-      apply inj_update;
-      intro;
-      rewrite > (eq_update_s_a_sa (update (mult_memory x y) 30 (mult_memory x y 30))
-       31 a);
-      rewrite > eq_update_s_a_sa;
-      reflexivity
-    ]
-  | cut (5 + 23 * S n = 5 + 23 * n + 23);
-    [ rewrite > Hcut; clear Hcut;
-      rewrite > breakpoint;
-      rewrite > H; clear H;
-      [2: apply le_S_S_to_le;
-        apply le_S;
-        apply H1
-      | cut (∃z.y-n=S z ∧ z < 255);
-         [ elim Hcut; clear Hcut;
-           elim H; clear H;
-           rewrite > H2;
-           (* instruction LDAd *)
-           change in ⊢ (? ? (? ? %) ?) with (3+20);
-           rewrite > breakpoint in ⊢ (? ? % ?);
-           whd in ⊢ (? ? (? % ?) ?);
-           normalize in ⊢ (? ? (? (? ? % ? ? ? ? ?) ?) ?);
-           change in ⊢ (? ? (? (? % ? ? ? ? ? ?) ?) ?)
-            with (byte_of_nat (S a));
-           change in ⊢ (? ? (? (? ? ? ? (? ? %) ? ? ?) ?) ?) with
-            (byte_of_nat (S a));
-           (* instruction BEQ *)
-           change in ⊢ (? ? (? ? %) ?) with (3+17);
-           rewrite > breakpoint in ⊢ (? ? % ?);
-           whd in ⊢ (? ? (? % ?) ?);
-           letin K ≝ (eq_eqbyte_x0_x0_byte_of_nat_S_false ? H3); clearbody K;
-           rewrite > K; clear K;
-           simplify in ⊢ (? ? (? (? ? % ? ? ? ? ?) ?) ?);
-           (* instruction LDAd *)
-           change in ⊢ (? ? (? ? %) ?) with (3+14);
-           rewrite > breakpoint in ⊢ (? ? % ?);
-           whd in ⊢ (? ? (? % ?) ?);
-           change in ⊢ (? ? (? (? % ? ? ? ? ? ?) ?) ?) with (byte_of_nat (x*n));
-           normalize in ⊢ (? ? (? (? ? % ? ? ? ? ?) ?) ?);
-           change in ⊢ (? ? (? (? ? ? ? % ? ? ?) ?) ?) with (eqbyte 〈x0, x0〉 (byte_of_nat (x*n)));
-           (* instruction DECd *)
-           change in ⊢ (? ? (? ? %) ?) with (5+9);
-           rewrite > breakpoint in ⊢ (? ? % ?);
-           whd in ⊢ (? ? (? % ?) ?);
-           change in ⊢ (? ? (? (? ? ? ? (? ? %) ? ? ?) ?) ?) with (bpred (byte_of_nat (S a)));
-           rewrite > (eq_bpred_S_a_a ? H3);
-           normalize in ⊢ (? ? (? (? ? % ? ? ? ? ?) ?) ?);
-           normalize in ⊢ (? ? (? (? ? ? ? ? ? (? ? % ?) ?) ?) ?);
-           cut (y - S n = a);
-            [2: rewrite > eq_minus_S_pred;
-                rewrite > H2;
-                reflexivity | ];
-           rewrite < Hcut; clear Hcut; clear H3; clear H2; clear a;          
-           (* instruction ADDd *)
-           change in ⊢ (? ? (? ? %) ?) with (3+6);
-           rewrite > breakpoint in ⊢ (? ? % ?);
-           whd in ⊢ (? ? (? % ?) ?);
-           change in ⊢ (? ? (? (? % ? ? ? ? ? ?) ?) ?) with
-            (plusbytenc (byte_of_nat (x*n)) x);
-           change in ⊢ (? ? (? (? ? ? ? (? ? %) ? ? ?) ?) ?) with
-            (plusbytenc (byte_of_nat (x*n)) x);
-           normalize in ⊢ (? ? (? (? ? % ? ? ? ? ?) ?) ?);
-           change in ⊢ (? ? (? (? ? ? ? ? % ? ?) ?) ?)
-            with (plusbytec (byte_of_nat (x*n)) x);
-           rewrite > plusbytenc_S;
-           (* instruction STAd *)
-           rewrite > (breakpoint ? 3 3);
-           whd in ⊢ (? ? (? % ?) ?);
-           normalize in ⊢ (? ? (? (? ? % ? ? ? ? ?) ?) ?);
-           (* instruction BRA *)
-           whd in ⊢ (? ? % ?);
-           normalize in ⊢ (? ? (? ? % ? ? ? ? ?) ?);
-           rewrite < pred_Sn;        
-           apply status_eq;
-            [1,2,3,4,7: normalize; reflexivity
-            | change with (plusbytec #(x*n) x = plusbytec #(x*n) x);
-              reflexivity
-            |6: intro;
-              simplify in ⊢ (? ? ? %);
-              normalize in ⊢ (? ? (? (? ? ? ? ? ? (? ? (? %) ?) ?) ?) ?);
-              change in ⊢ (? ? % ?) with
-               ((mult_memory x y){30↦x}{31↦#(S (y-S n))}{32↦#(x*n)}{31↦#(y-S n)}
-                {〈x2,x0〉↦ #(x*S n)} a);
-              apply inj_update;
-              intro;
-              apply inj_update;
-              intro;
-              rewrite > not_eq_a_b_to_eq_update_a_b; [2: apply H | ];
-              rewrite > not_eq_a_b_to_eq_update_a_b;
-               [ reflexivity
-               | assumption
-               ]
-            ]
-         | exists;
-            [ apply (y - S n)
-            | split;
-               [ rewrite < (minus_S_S y n);
-                 apply (minus_Sn_m (nat_of_byte y) (S n) H1)
-               | letin K ≝ (lt_nat_of_byte_256 y); clearbody K;
-                 letin K' ≝ (lt_minus_m y (S n) ? ?); clearbody K';
-                 [ apply (lt_to_le_to_lt O (S n) (nat_of_byte y) ? ?);
-                   autobatch
-                 | autobatch
-                 | autobatch
-                 ]
-               ]
-            ]
-         ]
-      ]
-    | rewrite > associative_plus;
-      rewrite < times_n_Sm;
-      rewrite > sym_plus in ⊢ (? ? ? (? ? %));
-      reflexivity
-    ] 
-  ]   
-qed.
-
-
-theorem test_x_y:
- ∀x,y:byte.
-  let i ≝ 14 + 23 * y in
-   execute (mult_status x y) i =
-    mk_status (#(x*y)) 20 0
-     (eqbyte 〈x0, x0〉 (#(x*y)))
-     (plusbytec (byte_of_nat (x*pred y)) x)
-     (update
-       (update (mult_memory x y) 31 〈x0, x0〉)
-       32 (byte_of_nat (x*y)))
-     0.
- intros;
- cut (14 + 23 * y = 5 + 23*y + 9);
-  [2: autobatch paramodulation;
-  | rewrite > Hcut; (* clear Hcut; *)
-    rewrite > (breakpoint (mult_status x y) (5 + 23*y) 9);
-    rewrite > loop_invariant';
-     [2: apply le_n
-     | rewrite < minus_n_n;
-       apply status_eq;
-        [1,2,3,4,5,7: normalize; reflexivity
-        | intro;
-          simplify in ⊢ (? ? ? %);
-          change in ⊢ (? ? % ?) with
-           (update (update (update (mult_memory x y) 30 x) 31 (byte_of_nat O)) 32
-(byte_of_nat (nat_of_byte x*nat_of_byte y)) a);
-          repeat (apply inj_update; intro);
-          apply (eq_update_s_a_sa ? 30)
-        ]
-     ]
-  ].
-qed.
\ No newline at end of file
diff --git a/helm/software/matita/library/assembly/vm.ma b/helm/software/matita/library/assembly/vm.ma
deleted file mode 100644 (file)
index 30d0731..0000000
+++ /dev/null
@@ -1,282 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||M||                                                             *)
-(*      ||A||       A project by Andrea Asperti                           *)
-(*      ||T||                                                             *)
-(*      ||I||       Developers:                                           *)
-(*      ||T||         The HELM team.                                      *)
-(*      ||A||         http://helm.cs.unibo.it                             *)
-(*      \   /                                                             *)
-(*       \ /        This file is distributed under the terms of the       *)
-(*        v         GNU General Public License Version 2                  *)
-(*                                                                        *)
-(**************************************************************************)
-
-set "baseuri" "cic:/matita/assembly/vm/".
-
-include "assembly/byte.ma".
-
-definition addr ≝ nat.
-
-definition addr_of_byte : byte → addr ≝ λb. nat_of_byte b.
-
-coercion cic:/matita/assembly/vm/addr_of_byte.con.
-
-inductive opcode: Type ≝
-   ADDd: opcode  (* 3 clock, 171 *)
- | BEQ: opcode   (* 3, 55 *)
- | BRA: opcode   (* 3, 48 *)
- | DECd: opcode  (* 5, 58 *)
- | LDAi: opcode  (* 2, 166 *)
- | LDAd: opcode  (* 3, 182 *)
- | STAd: opcode. (* 3, 183 *)
-
-let rec cycles_of_opcode op : nat ≝
- match op with
-  [ ADDd ⇒ 3
-  | BEQ ⇒ 3
-  | BRA ⇒ 3
-  | DECd ⇒ 5
-  | LDAi ⇒ 2
-  | LDAd ⇒ 3
-  | STAd ⇒ 3
-  ].
-
-definition opcodemap ≝
- [ couple ? ? ADDd (mk_byte xA xB);
-   couple ? ? BEQ (mk_byte x3 x7);
-   couple ? ? BRA (mk_byte x3 x0);
-   couple ? ? DECd (mk_byte x3 xA);
-   couple ? ? LDAi (mk_byte xA x6);
-   couple ? ? LDAd (mk_byte xB x6);
-   couple ? ? STAd (mk_byte xB x7) ].
-
-definition opcode_of_byte ≝
- λb.
-  let rec aux l ≝
-   match l with
-    [ nil ⇒ ADDd
-    | cons c tl ⇒
-       match c with
-        [ couple op n ⇒
-           match eqbyte n b with
-            [ true ⇒ op
-            | false ⇒ aux tl
-            ]]]
-  in
-   aux opcodemap.
-
-definition magic_of_opcode ≝
- λop1.
-  match op1 with
-   [ ADDd ⇒ 0
-   | BEQ ⇒  1 
-   | BRA ⇒  2
-   | DECd ⇒ 3
-   | LDAi ⇒ 4
-   | LDAd ⇒ 5
-   | STAd ⇒ 6 ].
-   
-definition opcodeeqb ≝
- λop1,op2. eqb (magic_of_opcode op1) (magic_of_opcode op2).
-
-definition byte_of_opcode : opcode → byte ≝
- λop.
-  let rec aux l ≝
-   match l with
-    [ nil ⇒ mk_byte x0 x0
-    | cons c tl ⇒
-       match c with
-        [ couple op' n ⇒
-           match opcodeeqb op op' with
-            [ true ⇒ n
-            | false ⇒ aux tl
-            ]]]
-  in
-   aux opcodemap.
-
-notation "hvbox(# break a)"
-  non associative with precedence 80
-for @{ 'byte_of_opcode $a }.
-interpretation "byte_of_opcode" 'byte_of_opcode a =
- (cic:/matita/assembly/vm/byte_of_opcode.con a).
-
-record status : Type ≝ {
-  acc : byte;
-  pc  : addr;
-  spc : addr;
-  zf  : bool;
-  cf  : bool;
-  mem : addr → byte;
-  clk : nat
-}.
-
-notation "{hvbox('Acc' ≝ acc ; break 'Pc' ≝ pc ; break 'Spc' ≝ spc ; break 'Fz' ≝ fz ; break 'Fc' ≝ fc ; break  'M' ≝ m ; break 'Clk' ≝ clk)}"
-non associative with precedence 80 for
- @{ 'mkstatus $acc $pc $spc $fz $fc $m $clk }.
-interpretation "mk_status" 'mkstatus acc pc spc fz fc m clk =
- (cic:/matita/assembly/vm/status.ind#xpointer(1/1/1) acc pc spc fz fc m clk).
-
-definition update ≝
- λf: addr → byte.λa.λv.λx.
-  match eqb x a with
-   [ true ⇒ v
-   | false ⇒ f x ].
-
-notation "hvbox(m break {a ↦ break v})" non associative with precedence 80 for 
- @{ 'update $m $a $v }.
-notation "hvbox(m break {a ↦ break v} \nbsp x)" non associative with precedence 80 for 
- @{ 'update4 $m $a $v $x }.
-interpretation "update" 'update m a v =
-  (cic:/matita/assembly/vm/update.con m a v).
-
-interpretation "update4" 'update4 m a v x =
-  (cic:/matita/assembly/vm/update.con m a v x).
-
-lemma update_update_a_a:
- ∀s,a,v1,v2,b.
-  update (update s a v1) a v2 b = update s a v2 b.
- intros;
- unfold update;
- unfold update;
- elim (eqb b a);
- reflexivity.
-qed. 
-
-lemma update_update_a_b:
- ∀s,a1,v1,a2,v2,b.
-  a1 ≠ a2 →
-   update (update s a1 v1) a2 v2 b = update (update s a2 v2) a1 v1 b.
- intros;
- unfold update;
- unfold update;
- apply (bool_elim ? (eqb b a1)); intros;
- apply (bool_elim ? (eqb b a2)); intros;
- simplify;
- [ elim H;
-   rewrite < (eqb_true_to_eq ? ? H1);
-   apply eqb_true_to_eq;
-   assumption
- |*: reflexivity
- ].
-qed.
-
-lemma eq_update_s_a_sa: ∀s,a,b. update s a (s a) b = s b.
- intros;
- unfold update;
- apply (bool_elim ? (eqb b a) ? ?); simplify; intros;
-  [ rewrite > (eqb_true_to_eq ? ? H);
-    reflexivity
-  | reflexivity
-  ]
-qed.
-
-lemma inj_update:
- ∀s,s',a,v,b. (a ≠ b → s b = s' b) → update s a v b = update s' a v b.
- intros;
- unfold update;
- apply (bool_elim ? (eqb b a) ? ?); simplify; intros;
-  [ reflexivity
-  | apply H;
-    intro;
-    autobatch
-  ]
-qed.
-
-lemma not_eq_a_b_to_eq_update_a_b: ∀s,a,b,v. a ≠ b → update s a v b = s b.
- intros;
- unfold update;
- rewrite > not_eq_to_eqb_false; simplify;
-  [ reflexivity
-  | intro;
-    autobatch
-  ]
-qed.
-
-definition mmod16 ≝ λn. nat_of_byte (byte_of_nat n).
-
-definition tick ≝
- λs:status. match s with [ mk_status acc pc spc zf cf mem clk ⇒
-  let opc ≝ opcode_of_byte (mem pc) in
-  let op1 ≝ mem (S pc) in
-  let clk' ≝ cycles_of_opcode opc in
-  match eqb (S clk) clk' with
-   [ true ⇒
-      match opc with
-       [ ADDd ⇒
-          let res ≝ plusbyte acc (mem op1) false in (* verify carrier! *)
-          let acc' ≝ match res with [ couple acc' _ ⇒ acc' ] in
-          let c' ≝ match res with [ couple _ c' ⇒ c'] in
-           mk_status acc' (2 + pc) spc
-            (eqbyte (mk_byte x0 x0) acc') c' mem 0 (* verify carrier! *)
-       | BEQ ⇒
-          mk_status
-           acc
-           (match zf with
-             [ true ⇒ mmod16 (2 + op1 + pc) (*\mod 256*)   (* signed!!! *)
-                      (* FIXME: can't work - address truncated to 8 bits *)
-             | false ⇒ 2 + pc
-             ])
-           spc
-           zf
-           cf
-           mem
-           0
-       | BRA ⇒
-          mk_status
-           acc (mmod16 (2 + op1 + pc) (*\mod 256*)) (* signed!!! *)
-                                      (* FIXME: same as above *)
-           spc
-           zf
-           cf
-           mem
-           0
-       | DECd ⇒
-          let x ≝ bpred (mem op1) in (* signed!!! *)
-          let mem' ≝ update mem op1 x in
-           mk_status acc (2 + pc) spc
-            (eqbyte (mk_byte x0 x0) x) cf mem' 0 (* check zb!!! *)
-       | LDAi ⇒
-          mk_status op1 (2 + pc) spc (eqbyte (mk_byte x0 x0) op1) cf mem 0
-       | LDAd ⇒
-          let x ≝ mem op1 in
-           mk_status x (2 + pc) spc (eqbyte (mk_byte x0 x0) x) cf mem 0
-       | STAd ⇒
-          mk_status acc (2 + pc) spc zf cf
-           (update mem op1 acc) 0
-       ]
-   | false ⇒
-       mk_status
-        acc pc spc zf cf mem (S clk)
-   ]].
-
-let rec execute s n on n ≝
- match n with
-  [ O ⇒ s
-  | S n' ⇒ execute (tick s) n'
-  ].
-  
-lemma breakpoint:
- ∀s,n1,n2. execute s (n1 + n2) = execute (execute s n1) n2.
- intros;
- generalize in match s; clear s;
- elim n1;
-  [ reflexivity
-  | simplify;
-    apply H;
-  ]
-qed.
-
-axiom status_eq:
- ∀s,s'.
-  acc s = acc s' →
-  pc s = pc s' →
-  spc s = spc s' →
-  zf s = zf s' →
-  cf s = cf s' →
-  (∀a. mem s a = mem s' a) →
-  clk s = clk s' →
-   s=s'.
index 0e508843088b36521547410b00e888c2e46c3431..de935c32b28a54d28724f77e0e6a7215c4b986dd 100644 (file)
-Z/dirichlet_product.ma Z/sigma_p.ma Z/times.ma nat/primes.ma
-Z/moebius.ma Z/sigma_p.ma nat/factorization.ma
-Z/times.ma Z/plus.ma nat/lt_arith.ma
-Z/orders.ma Z/z.ma logic/connectives.ma nat/orders.ma
-Z/inversion.ma Z/dirichlet_product.ma Z/moebius.ma
-Z/plus.ma Z/z.ma nat/minus.ma
-Z/compare.ma Z/orders.ma nat/compare.ma
-Z/sigma_p.ma Z/plus.ma Z/times.ma nat/generic_iter_p.ma nat/ord.ma nat/primes.ma
-Z/z.ma datatypes/bool.ma nat/nat.ma
-assembly/test.ma assembly/vm.ma
-assembly/byte.ma assembly/exadecimal.ma
-assembly/vm.ma assembly/byte.ma
-assembly/exadecimal.ma assembly/extra.ma
-assembly/extra.ma list/list.ma nat/div_and_mod.ma nat/primes.ma
-datatypes/constructors.ma logic/equality.ma
+why_for_assembly.ma freescale/multivm.ma
+freescale/memory_func.ma freescale/memory_struct.ma
+freescale/memory_trees.ma freescale/memory_struct.ma
+freescale/medium_tests.ma freescale/medium_tests_tools.ma
+freescale/table_HC05.ma freescale/opcode.ma
+freescale/table_HC08.ma freescale/opcode.ma
+freescale/status.ma freescale/memory_abs.ma
+freescale/aux_bases.ma freescale/word16.ma
+freescale/word16.ma freescale/byte8.ma
+freescale/memory_struct.ma freescale/translation.ma
+freescale/memory_abs.ma freescale/memory_bits.ma freescale/memory_func.ma freescale/memory_trees.ma
+freescale/opcode.ma freescale/aux_bases.ma
+freescale/memory_bits.ma freescale/memory_trees.ma
+freescale/table_RS08.ma freescale/opcode.ma
+freescale/table_HCS08.ma freescale/opcode.ma
+freescale/load_write.ma freescale/model.ma
+freescale/medium_tests_tools.ma freescale/multivm.ma
+freescale/byte8.ma freescale/exadecim.ma
+freescale/exadecim.ma freescale/extra.ma
+freescale/extra.ma datatypes/constructors.ma list/list.ma logic/connectives.ma nat/div_and_mod.ma nat/primes.ma
+freescale/multivm.ma freescale/load_write.ma
+freescale/micro_tests.ma freescale/multivm.ma
+freescale/model.ma freescale/status.ma
+freescale/translation.ma freescale/table_HC05.ma freescale/table_HC08.ma freescale/table_HCS08.ma freescale/table_RS08.ma
 datatypes/compare.ma 
+datatypes/constructors.ma logic/equality.ma
 datatypes/bool.ma higher_order_defs/functions.ma logic/equality.ma
-algebra/groups.ma algebra/monoids.ma datatypes/bool.ma logic/connectives.ma nat/compare.ma nat/le_arith.ma
+logic/coimplication.ma logic/connectives.ma
+logic/equality.ma higher_order_defs/relations.ma logic/connectives.ma
+logic/connectives.ma 
+logic/connectives2.ma higher_order_defs/relations.ma
+algebra/monoids.ma algebra/semigroups.ma
 algebra/finite_groups.ma algebra/groups.ma nat/relevant_equations.ma
 algebra/semigroups.ma higher_order_defs/functions.ma
-algebra/monoids.ma algebra/semigroups.ma
-algebra/CoRN/SetoidInc.ma algebra/CoRN/SetoidFun.ma
-algebra/CoRN/SemiGroups.ma algebra/CoRN/SetoidInc.ma
+algebra/groups.ma algebra/monoids.ma datatypes/bool.ma logic/connectives.ma nat/compare.ma nat/le_arith.ma
 algebra/CoRN/Setoids.ma Z/plus.ma datatypes/constructors.ma higher_order_defs/relations.ma logic/equality.ma nat/nat.ma
+algebra/CoRN/SemiGroups.ma algebra/CoRN/SetoidInc.ma
 algebra/CoRN/SetoidFun.ma algebra/CoRN/Setoids.ma higher_order_defs/relations.ma
-demo/propositional_sequent_calculus.ma datatypes/constructors.ma list/sort.ma nat/compare.ma nat/plus.ma
-demo/power_derivative.ma nat/compare.ma nat/nat.ma nat/orders.ma nat/plus.ma
+algebra/CoRN/SetoidInc.ma algebra/CoRN/SetoidFun.ma
 list/sort.ma datatypes/bool.ma datatypes/constructors.ma list/list.ma
 list/list.ma higher_order_defs/functions.ma logic/equality.ma nat/nat.ma
-logic/equality.ma higher_order_defs/relations.ma logic/connectives.ma
-logic/connectives.ma 
-logic/coimplication.ma logic/connectives.ma
-logic/connectives2.ma higher_order_defs/relations.ma
-nat/div_and_mod.ma datatypes/constructors.ma nat/minus.ma
-nat/count.ma nat/permutation.ma nat/relevant_equations.ma nat/sigma_and_pi.ma
-nat/factorial2.ma nat/exp.ma nat/factorial.ma
-nat/totient1.ma nat/compare.ma nat/gcd_properties1.ma nat/iteration2.ma nat/totient.ma
+Z/moebius.ma Z/sigma_p.ma nat/factorization.ma
+Z/plus.ma Z/z.ma nat/minus.ma
+Z/compare.ma Z/orders.ma nat/compare.ma
+Z/dirichlet_product.ma Z/sigma_p.ma Z/times.ma nat/primes.ma
+Z/z.ma datatypes/bool.ma nat/nat.ma
+Z/inversion.ma Z/dirichlet_product.ma Z/moebius.ma
+Z/orders.ma Z/z.ma logic/connectives.ma nat/orders.ma
+Z/sigma_p.ma Z/plus.ma Z/times.ma nat/generic_iter_p.ma nat/ord.ma nat/primes.ma
+Z/times.ma Z/plus.ma nat/lt_arith.ma
+technicalities/setoids.ma datatypes/constructors.ma logic/coimplication.ma logic/connectives2.ma
+demo/power_derivative.ma nat/compare.ma nat/nat.ma nat/orders.ma nat/plus.ma
+demo/propositional_sequent_calculus.ma datatypes/constructors.ma list/sort.ma nat/compare.ma nat/plus.ma
+decidable_kit/eqtype.ma datatypes/constructors.ma decidable_kit/decidable.ma
+decidable_kit/fgraph.ma decidable_kit/fintype.ma
+decidable_kit/list_aux.ma decidable_kit/eqtype.ma list/list.ma nat/plus.ma
+decidable_kit/streicher.ma logic/connectives.ma logic/equality.ma
+decidable_kit/decidable.ma datatypes/bool.ma decidable_kit/streicher.ma logic/connectives.ma nat/compare.ma
+decidable_kit/fintype.ma decidable_kit/eqtype.ma decidable_kit/list_aux.ma
+nat/plus.ma nat/nat.ma
+nat/ord.ma datatypes/constructors.ma nat/exp.ma nat/gcd.ma nat/nth_prime.ma nat/relevant_equations.ma
+nat/totient.ma nat/chinese_reminder.ma nat/iteration2.ma
+nat/pi_p.ma nat/generic_iter_p.ma nat/iteration2.ma nat/primes.ma
+nat/compare.ma datatypes/bool.ma datatypes/compare.ma nat/orders.ma
+nat/le_arith.ma nat/orders.ma nat/times.ma
 nat/congruence.ma nat/primes.ma nat/relevant_equations.ma
+nat/count.ma nat/permutation.ma nat/relevant_equations.ma nat/sigma_and_pi.ma
+nat/gcd_properties1.ma nat/gcd.ma
 nat/minus.ma nat/compare.ma nat/le_arith.ma
-nat/chebyshev.ma nat/factorial2.ma nat/factorization.ma nat/log.ma nat/pi_p.ma
-nat/chinese_reminder.ma nat/congruence.ma nat/exp.ma nat/gcd.ma nat/permutation.ma
-nat/permutation.ma nat/compare.ma nat/sigma_and_pi.ma
-nat/factorial.ma nat/le_arith.ma
-nat/nth_prime.ma nat/lt_arith.ma nat/primes.ma
-nat/binomial.ma nat/factorial2.ma nat/iteration2.ma
-nat/fermat_little_theorem.ma nat/congruence.ma nat/exp.ma nat/gcd.ma nat/permutation.ma
 nat/exp.ma nat/div_and_mod.ma nat/lt_arith.ma
+nat/div_and_mod.ma datatypes/constructors.ma nat/minus.ma
 nat/minimization.ma nat/minus.ma
-nat/gcd.ma nat/lt_arith.ma nat/primes.ma
-nat/chebyshev_thm.ma nat/neper.ma
-nat/neper.ma nat/binomial.ma nat/chebyshev.ma nat/div_and_mod_diseq.ma nat/iteration2.ma nat/log.ma
+nat/relevant_equations.ma nat/gcd.ma nat/minus.ma nat/times.ma
 nat/sigma_and_pi.ma nat/exp.ma nat/factorial.ma nat/lt_arith.ma
-nat/le_arith.ma nat/orders.ma nat/times.ma
-nat/times.ma nat/plus.ma
-nat/div_and_mod_diseq.ma nat/lt_arith.ma
+nat/fermat_little_theorem.ma nat/congruence.ma nat/exp.ma nat/gcd.ma nat/permutation.ma
 nat/generic_iter_p.ma nat/div_and_mod_diseq.ma nat/ord.ma nat/primes.ma
-nat/gcd_properties1.ma nat/gcd.ma
-nat/totient.ma nat/chinese_reminder.ma nat/iteration2.ma
-nat/lt_arith.ma nat/div_and_mod.ma
-nat/factorization.ma nat/ord.ma
-nat/log.ma datatypes/constructors.ma nat/div_and_mod_diseq.ma nat/iteration2.ma nat/minimization.ma nat/primes.ma nat/relevant_equations.ma
-nat/iteration2.ma nat/count.ma nat/generic_iter_p.ma nat/ord.ma nat/primes.ma
 nat/nat.ma higher_order_defs/functions.ma
-nat/relevant_equations.ma nat/gcd.ma nat/minus.ma nat/times.ma
-nat/map_iter_p.ma nat/count.ma nat/permutation.ma
+nat/nth_prime.ma nat/lt_arith.ma nat/primes.ma
+nat/gcd.ma nat/lt_arith.ma nat/primes.ma
 nat/orders.ma higher_order_defs/ordering.ma logic/connectives.ma nat/nat.ma
-nat/pi_p.ma nat/generic_iter_p.ma nat/iteration2.ma nat/primes.ma
-nat/plus.ma nat/nat.ma
-nat/euler_theorem.ma nat/nat.ma nat/map_iter_p.ma nat/totient.ma
-nat/compare.ma datatypes/bool.ma datatypes/compare.ma nat/orders.ma
 nat/primes.ma logic/connectives.ma nat/div_and_mod.ma nat/factorial.ma nat/minimization.ma nat/sigma_and_pi.ma
-nat/ord.ma datatypes/constructors.ma nat/exp.ma nat/gcd.ma nat/nth_prime.ma nat/relevant_equations.ma
-Q/Qaxioms.ma Z/compare.ma Z/times.ma nat/iteration2.ma
+nat/euler_theorem.ma nat/nat.ma nat/map_iter_p.ma nat/totient.ma
+nat/chinese_reminder.ma nat/congruence.ma nat/exp.ma nat/gcd.ma nat/permutation.ma
+nat/factorial.ma nat/le_arith.ma
+nat/lt_arith.ma nat/div_and_mod.ma
+nat/permutation.ma nat/compare.ma nat/sigma_and_pi.ma
+nat/totient1.ma nat/compare.ma nat/gcd_properties1.ma nat/iteration2.ma nat/totient.ma
+nat/factorization.ma nat/ord.ma
+nat/map_iter_p.ma nat/count.ma nat/permutation.ma
+nat/iteration2.ma nat/count.ma nat/generic_iter_p.ma nat/ord.ma nat/primes.ma
+nat/times.ma nat/plus.ma
+nat/factorial2.ma nat/exp.ma nat/factorial.ma
+nat/chebyshev.ma nat/factorial2.ma nat/factorization.ma nat/log.ma nat/pi_p.ma
+nat/log.ma datatypes/constructors.ma nat/div_and_mod_diseq.ma nat/iteration2.ma nat/minimization.ma nat/primes.ma nat/relevant_equations.ma
+nat/div_and_mod_diseq.ma nat/lt_arith.ma
+nat/binomial.ma nat/factorial2.ma nat/iteration2.ma
+nat/neper.ma nat/binomial.ma nat/chebyshev.ma nat/div_and_mod_diseq.ma nat/iteration2.ma nat/log.ma
+nat/chebyshev_thm.ma nat/neper.ma
+higher_order_defs/functions.ma logic/equality.ma
+higher_order_defs/ordering.ma logic/equality.ma
+higher_order_defs/relations.ma logic/connectives.ma
 Q/q.ma Z/compare.ma Z/plus.ma
-technicalities/setoids.ma datatypes/constructors.ma logic/coimplication.ma logic/connectives2.ma
+Q/Qaxioms.ma Z/compare.ma Z/times.ma nat/iteration2.ma
 Fsub/part1a.ma Fsub/defn.ma
-Fsub/util.ma list/list.ma logic/equality.ma nat/compare.ma
 Fsub/defn.ma Fsub/util.ma
+Fsub/util.ma list/list.ma logic/equality.ma nat/compare.ma
 Fsub/part1a_inversion.ma Fsub/defn.ma
-decidable_kit/streicher.ma logic/connectives.ma logic/equality.ma
-decidable_kit/decidable.ma datatypes/bool.ma decidable_kit/streicher.ma logic/connectives.ma nat/compare.ma
-decidable_kit/fgraph.ma decidable_kit/fintype.ma
-decidable_kit/eqtype.ma datatypes/constructors.ma decidable_kit/decidable.ma
-decidable_kit/list_aux.ma decidable_kit/eqtype.ma list/list.ma nat/plus.ma
-decidable_kit/fintype.ma decidable_kit/eqtype.ma decidable_kit/list_aux.ma
-higher_order_defs/relations.ma logic/connectives.ma
-higher_order_defs/functions.ma logic/equality.ma
-higher_order_defs/ordering.ma logic/equality.ma
diff --git a/helm/software/matita/library/freescale/aux_bases.ma b/helm/software/matita/library/freescale/aux_bases.ma
new file mode 100644 (file)
index 0000000..c2aa8aa
--- /dev/null
@@ -0,0 +1,117 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/aux_bases/".
+
+(*include "/media/VIRTUOSO/freescale/word16.ma".*)
+include "freescale/word16.ma".
+
+(* ************************ *)
+(* DEFINIZIONE DEGLI OTTALI *)
+(* ************************ *)
+
+inductive oct : Type ≝
+  o0: oct
+| o1: oct
+| o2: oct
+| o3: oct
+| o4: oct
+| o5: oct
+| o6: oct
+| o7: oct.
+
+(* ottali → esadecimali *)
+definition exadecim_of_oct ≝
+λo:oct.
+ match o with
+  [ o0 ⇒ x0 | o1 ⇒ x1 | o2 ⇒ x2 | o3 ⇒ x3
+  | o4 ⇒ x4 | o5 ⇒ x5 | o6 ⇒ x6 | o7 ⇒ x7 ].
+
+coercion cic:/matita/freescale/aux_bases/exadecim_of_oct.con.
+
+(* iteratore sugli ottali *)
+definition forall_oct ≝ λP.
+ P o0 ⊗ P o1 ⊗ P o2 ⊗ P o3 ⊗ P o4 ⊗ P o5 ⊗ P o6 ⊗ P o7.
+
+(* ***************************** *)
+(* DEFINIZIONE DEI BITRIGESIMALI *)
+(* ***************************** *)
+
+inductive bitrigesim : Type ≝
+  t00: bitrigesim
+| t01: bitrigesim
+| t02: bitrigesim
+| t03: bitrigesim
+| t04: bitrigesim
+| t05: bitrigesim
+| t06: bitrigesim
+| t07: bitrigesim
+| t08: bitrigesim
+| t09: bitrigesim
+| t0A: bitrigesim
+| t0B: bitrigesim
+| t0C: bitrigesim
+| t0D: bitrigesim
+| t0E: bitrigesim
+| t0F: bitrigesim
+| t10: bitrigesim
+| t11: bitrigesim
+| t12: bitrigesim
+| t13: bitrigesim
+| t14: bitrigesim
+| t15: bitrigesim
+| t16: bitrigesim
+| t17: bitrigesim
+| t18: bitrigesim
+| t19: bitrigesim
+| t1A: bitrigesim
+| t1B: bitrigesim
+| t1C: bitrigesim
+| t1D: bitrigesim
+| t1E: bitrigesim
+| t1F: bitrigesim.
+
+(* bitrigesimali → byte *)
+definition byte8_of_bitrigesim ≝
+λt:bitrigesim.
+ match t with
+  [ t00 ⇒ 〈x0,x0〉 | t01 ⇒ 〈x0,x1〉 | t02 ⇒ 〈x0,x2〉 | t03 ⇒ 〈x0,x3〉
+  | t04 ⇒ 〈x0,x4〉 | t05 ⇒ 〈x0,x5〉 | t06 ⇒ 〈x0,x6〉 | t07 ⇒ 〈x0,x7〉
+  | t08 ⇒ 〈x0,x8〉 | t09 ⇒ 〈x0,x9〉 | t0A ⇒ 〈x0,xA〉 | t0B ⇒ 〈x0,xB〉
+  | t0C ⇒ 〈x0,xC〉 | t0D ⇒ 〈x0,xD〉 | t0E ⇒ 〈x0,xE〉 | t0F ⇒ 〈x0,xF〉
+  | t10 ⇒ 〈x1,x0〉 | t11 ⇒ 〈x1,x1〉 | t12 ⇒ 〈x1,x2〉 | t13 ⇒ 〈x1,x3〉
+  | t14 ⇒ 〈x1,x4〉 | t15 ⇒ 〈x1,x5〉 | t16 ⇒ 〈x1,x6〉 | t17 ⇒ 〈x1,x7〉
+  | t18 ⇒ 〈x1,x8〉 | t19 ⇒ 〈x1,x9〉 | t1A ⇒ 〈x1,xA〉 | t1B ⇒ 〈x1,xB〉
+  | t1C ⇒ 〈x1,xC〉 | t1D ⇒ 〈x1,xD〉 | t1E ⇒ 〈x1,xE〉 | t1F ⇒ 〈x1,xF〉 ].
+
+coercion cic:/matita/freescale/aux_bases/byte8_of_bitrigesim.con.
+
+(* iteratore sui bitrigesimali *)
+definition forall_bitrigesim ≝ λP.
+ P t00 ⊗ P t01 ⊗ P t02 ⊗ P t03 ⊗ P t04 ⊗ P t05 ⊗ P t06 ⊗ P t07 ⊗
+ P t08 ⊗ P t09 ⊗ P t0A ⊗ P t0B ⊗ P t0C ⊗ P t0D ⊗ P t0E ⊗ P t0F ⊗
+ P t10 ⊗ P t11 ⊗ P t12 ⊗ P t13 ⊗ P t14 ⊗ P t15 ⊗ P t16 ⊗ P t17 ⊗
+ P t18 ⊗ P t19 ⊗ P t1A ⊗ P t1B ⊗ P t1C ⊗ P t1D ⊗ P t1E ⊗ P t1F.
diff --git a/helm/software/matita/library/freescale/byte8.ma b/helm/software/matita/library/freescale/byte8.ma
new file mode 100644 (file)
index 0000000..fedc5a6
--- /dev/null
@@ -0,0 +1,556 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/byte8".
+
+(*include "/media/VIRTUOSO/freescale/exadecim.ma".*)
+include "freescale/exadecim.ma".
+
+(* ******************** *)
+(* DEFINIZIONE DEI BYTE *)
+(* ******************** *)
+
+record byte8 : Type ≝
+ {
+ b8h: exadecim;
+ b8l: exadecim
+ }.
+
+(* \langle \rangle *)
+notation "〈x,y〉" non associative with precedence 80
+ for @{ 'mk_byte8 $x $y }.
+interpretation "mk_byte8" 'mk_byte8 x y = 
+ (cic:/matita/freescale/byte8/byte8.ind#xpointer(1/1/1) x y).
+
+(* operatore = *)
+definition eq_b8 ≝ λb1,b2:byte8.(eq_ex (b8h b1) (b8h b2)) ⊗ (eq_ex (b8l b1) (b8l b2)).
+
+(* operatore < *)
+definition lt_b8 ≝
+λb1,b2:byte8.match lt_ex (b8h b1) (b8h b2) with
+ [ true ⇒ true
+ | false ⇒ match gt_ex (b8h b1) (b8h b2) with
+  [ true ⇒ false
+  | false ⇒ lt_ex (b8l b1) (b8l b2) ]].
+
+(* operatore ≤ *)
+definition le_b8 ≝ λb1,b2:byte8.(eq_b8 b1 b2) ⊕ (lt_b8 b1 b2). 
+
+(* operatore > *)
+definition gt_b8 ≝ λb1,b2:byte8.⊖ (le_b8 b1 b2).
+
+(* operatore ≥ *)
+definition ge_b8 ≝ λb1,b2:byte8.⊖ (lt_b8 b1 b2).
+
+(* operatore and *)
+definition and_b8 ≝
+λb1,b2:byte8.mk_byte8 (and_ex (b8h b1) (b8h b2)) (and_ex (b8l b1) (b8l b2)).
+
+(* operatore or *)
+definition or_b8 ≝
+λb1,b2:byte8.mk_byte8 (or_ex (b8h b1) (b8h b2)) (or_ex (b8l b1) (b8l b2)).
+
+(* operatore xor *)
+definition xor_b8 ≝
+λb1,b2:byte8.mk_byte8 (xor_ex (b8h b1) (b8h b2)) (xor_ex (b8l b1) (b8l b2)).
+
+(* operatore rotazione destra con carry *)
+definition rcr_b8 ≝
+λb:byte8.λc:bool.match rcr_ex (b8h b) c with
+ [ pairT bh' c' ⇒ match rcr_ex (b8l b) c' with
+  [ pairT bl' c'' ⇒ pairT ?? (mk_byte8 bh' bl') c'' ]]. 
+
+(* operatore shift destro *)
+definition shr_b8 ≝
+λb:byte8.match rcr_ex (b8h b) false with
+ [ pairT bh' c' ⇒ match rcr_ex (b8l b) c' with
+  [ pairT bl' c'' ⇒ pairT ?? (mk_byte8 bh' bl') c'' ]].
+
+(* operatore rotazione destra *)
+definition ror_b8 ≝
+λb:byte8.match rcr_ex (b8h b) false with
+ [ pairT bh' c' ⇒ match rcr_ex (b8l b) c' with
+  [ pairT bl' c'' ⇒ match c'' with
+   [ true ⇒ mk_byte8 (or_ex x8 bh') bl'
+   | false ⇒ mk_byte8 bh' bl' ]]].
+
+(* operatore rotazione destra n-volte *)
+let rec ror_b8_n (b:byte8) (n:nat) on n ≝
+ match n with
+  [ O ⇒ b
+  | S n' ⇒ ror_b8_n (ror_b8 b) n' ].
+
+(* operatore rotazione sinistra con carry *)
+definition rcl_b8 ≝
+λb:byte8.λc:bool.match rcl_ex (b8l b) c with
+ [ pairT bl' c' ⇒ match rcl_ex (b8h b) c' with
+  [ pairT bh' c'' ⇒ pairT ?? (mk_byte8 bh' bl') c'' ]]. 
+
+(* operatore shift sinistro *)
+definition shl_b8 ≝
+λb:byte8.match rcl_ex (b8l b) false with
+ [ pairT bl' c' ⇒ match rcl_ex (b8h b) c' with
+  [ pairT bh' c'' ⇒ pairT ?? (mk_byte8 bh' bl') c'' ]].
+
+(* operatore rotazione sinistra *)
+definition rol_b8 ≝
+λb:byte8.match rcl_ex (b8l b) false with
+ [ pairT bl' c' ⇒ match rcl_ex (b8h b) c' with
+  [ pairT bh' c'' ⇒ match c'' with
+   [ true ⇒ mk_byte8 bh' (or_ex x1 bl')
+   | false ⇒ mk_byte8 bh' bl' ]]].
+
+(* operatore rotazione sinistra n-volte *)
+let rec rol_b8_n (b:byte8) (n:nat) on n ≝
+ match n with
+  [ O ⇒ b
+  | S n' ⇒ rol_b8_n (rol_b8 b) n' ].
+
+(* operatore not/complemento a 1 *)
+definition not_b8 ≝
+λb:byte8.mk_byte8 (not_ex (b8h b)) (not_ex (b8l b)).
+
+(* operatore somma con carry *)
+definition plus_b8 ≝
+λb1,b2:byte8.λc:bool.
+ match plus_ex (b8l b1) (b8l b2) c with
+  [ pairT l c' ⇒ match plus_ex (b8h b1) (b8h b2) c' with
+   [ pairT h c'' ⇒ pairT ?? (mk_byte8 h l) c'' ]].
+
+(* operatore somma senza carry *)
+definition plus_b8nc ≝
+λb1,b2:byte8.fstT ?? (plus_b8 b1 b2 false).
+
+(* operatore carry della somma *)
+definition plus_b8c ≝
+λb1,b2:byte8.sndT ?? (plus_b8 b1 b2 false).
+
+(* operatore Most Significant Bit *)
+definition MSB_b8 ≝ λb:byte8.eq_ex x8 (and_ex x8 (b8h b)).
+
+(* byte → naturali *)
+definition nat_of_byte8 ≝ λb:byte8.16*(b8h b) + (b8l b).
+
+coercion cic:/matita/freescale/byte8/nat_of_byte8.con.
+
+(* naturali → byte *)
+definition byte8_of_nat ≝ λn.mk_byte8 (exadecim_of_nat (n/16)) (exadecim_of_nat n).
+
+(* operatore predecessore *)
+definition pred_b8 ≝
+λb:byte8.match eq_ex (b8l b) x0 with
+ [ true ⇒ mk_byte8 (pred_ex (b8h b)) (pred_ex (b8l b))
+ | false ⇒ mk_byte8 (b8h b) (pred_ex (b8l b)) ]. 
+
+(* operatore successore *)
+definition succ_b8 ≝
+λb:byte8.match eq_ex (b8l b) xF with
+ [ true ⇒ mk_byte8 (succ_ex (b8h b)) (succ_ex (b8l b))
+ | false ⇒ mk_byte8 (b8h b) (succ_ex (b8l b)) ]. 
+
+(* operatore neg/complemento a 2 *)
+definition compl_b8 ≝
+λb:byte8.match MSB_b8 b with
+ [ true ⇒ succ_b8 (not_b8 b)
+ | false ⇒ not_b8 (pred_b8 b) ].
+
+(* operatore moltiplicazione senza segno: e*e=[0x00,0xE1] *)
+definition mul_ex ≝
+λe1,e2:exadecim.match e1 with
+ [ x0 ⇒ match e2 with
+  [ x0 ⇒ 〈x0,x0〉   | x1 ⇒ 〈x0,x0〉   | x2 ⇒ 〈x0,x0〉   | x3 ⇒ 〈x0,x0〉
+  | x4 ⇒ 〈x0,x0〉   | x5 ⇒ 〈x0,x0〉   | x6 ⇒ 〈x0,x0〉   | x7 ⇒ 〈x0,x0〉
+  | x8 ⇒ 〈x0,x0〉   | x9 ⇒ 〈x0,x0〉   | xA ⇒ 〈x0,x0〉   | xB ⇒ 〈x0,x0〉
+  | xC ⇒ 〈x0,x0〉   | xD ⇒ 〈x0,x0〉   | xE ⇒ 〈x0,x0〉   | xF ⇒ 〈x0,x0〉 ]
+ | x1 ⇒ match e2 with
+  [ x0 ⇒ 〈x0,x0〉   | x1 ⇒ 〈x0,x1〉   | x2 ⇒ 〈x0,x2〉   | x3 ⇒ 〈x0,x3〉
+  | x4 ⇒ 〈x0,x4〉   | x5 ⇒ 〈x0,x5〉   | x6 ⇒ 〈x0,x6〉   | x7 ⇒ 〈x0,x7〉
+  | x8 ⇒ 〈x0,x8〉   | x9 ⇒ 〈x0,x9〉   | xA ⇒ 〈x0,xA〉   | xB ⇒ 〈x0,xB〉
+  | xC ⇒ 〈x0,xC〉   | xD ⇒ 〈x0,xD〉   | xE ⇒ 〈x0,xE〉   | xF ⇒ 〈x0,xF〉 ]
+ | x2 ⇒ match e2 with
+  [ x0 ⇒ 〈x0,x0〉   | x1 ⇒ 〈x0,x2〉   | x2 ⇒ 〈x0,x4〉   | x3 ⇒ 〈x0,x6〉
+  | x4 ⇒ 〈x0,x8〉   | x5 ⇒ 〈x0,xA〉   | x6 ⇒ 〈x0,xC〉   | x7 ⇒ 〈x0,xE〉
+  | x8 ⇒ 〈x1,x0〉   | x9 ⇒ 〈x1,x2〉   | xA ⇒ 〈x1,x4〉   | xB ⇒ 〈x1,x6〉
+  | xC ⇒ 〈x1,x8〉   | xD ⇒ 〈x1,xA〉   | xE ⇒ 〈x1,xC〉   | xF ⇒ 〈x1,xE〉 ]
+ | x3 ⇒ match e2 with
+  [ x0 ⇒ 〈x0,x0〉   | x1 ⇒ 〈x0,x3〉   | x2 ⇒ 〈x0,x6〉   | x3 ⇒ 〈x0,x9〉
+  | x4 ⇒ 〈x0,xC〉   | x5 ⇒ 〈x0,xF〉   | x6 ⇒ 〈x1,x2〉   | x7 ⇒ 〈x1,x5〉
+  | x8 ⇒ 〈x1,x8〉   | x9 ⇒ 〈x1,xB〉   | xA ⇒ 〈x1,xE〉   | xB ⇒ 〈x2,x1〉
+  | xC ⇒ 〈x2,x4〉   | xD ⇒ 〈x2,x7〉   | xE ⇒ 〈x2,xA〉   | xF ⇒ 〈x2,xD〉 ]
+ | x4 ⇒ match e2 with
+  [ x0 ⇒ 〈x0,x0〉   | x1 ⇒ 〈x0,x4〉   | x2 ⇒ 〈x0,x8〉   | x3 ⇒ 〈x0,xC〉
+  | x4 ⇒ 〈x1,x0〉   | x5 ⇒ 〈x1,x4〉   | x6 ⇒ 〈x1,x8〉   | x7 ⇒ 〈x1,xC〉
+  | x8 ⇒ 〈x2,x0〉   | x9 ⇒ 〈x2,x4〉   | xA ⇒ 〈x2,x8〉   | xB ⇒ 〈x2,xC〉
+  | xC ⇒ 〈x3,x0〉   | xD ⇒ 〈x3,x4〉   | xE ⇒ 〈x3,x8〉   | xF ⇒ 〈x3,xC〉 ]
+ | x5 ⇒ match e2 with
+  [ x0 ⇒ 〈x0,x0〉   | x1 ⇒ 〈x0,x5〉   | x2 ⇒ 〈x0,xA〉   | x3 ⇒ 〈x0,xF〉
+  | x4 ⇒ 〈x1,x4〉   | x5 ⇒ 〈x1,x9〉   | x6 ⇒ 〈x1,xE〉   | x7 ⇒ 〈x2,x3〉
+  | x8 ⇒ 〈x2,x8〉   | x9 ⇒ 〈x2,xD〉   | xA ⇒ 〈x3,x2〉   | xB ⇒ 〈x3,x7〉
+  | xC ⇒ 〈x3,xC〉   | xD ⇒ 〈x4,x1〉   | xE ⇒ 〈x4,x6〉   | xF ⇒ 〈x4,xB〉 ]
+ | x6 ⇒ match e2 with
+  [ x0 ⇒ 〈x0,x0〉   | x1 ⇒ 〈x0,x6〉   | x2 ⇒ 〈x0,xC〉   | x3 ⇒ 〈x1,x2〉
+  | x4 ⇒ 〈x1,x8〉   | x5 ⇒ 〈x1,xE〉   | x6 ⇒ 〈x2,x4〉   | x7 ⇒ 〈x2,xA〉
+  | x8 ⇒ 〈x3,x0〉   | x9 ⇒ 〈x3,x6〉   | xA ⇒ 〈x3,xC〉   | xB ⇒ 〈x4,x2〉
+  | xC ⇒ 〈x4,x8〉   | xD ⇒ 〈x4,xE〉   | xE ⇒ 〈x5,x4〉   | xF ⇒ 〈x5,xA〉 ]
+ | x7 ⇒ match e2 with
+  [ x0 ⇒ 〈x0,x0〉   | x1 ⇒ 〈x0,x7〉   | x2 ⇒ 〈x0,xE〉   | x3 ⇒ 〈x1,x5〉
+  | x4 ⇒ 〈x1,xC〉   | x5 ⇒ 〈x2,x3〉   | x6 ⇒ 〈x2,xA〉   | x7 ⇒ 〈x3,x1〉
+  | x8 ⇒ 〈x3,x8〉   | x9 ⇒ 〈x3,xF〉   | xA ⇒ 〈x4,x6〉   | xB ⇒ 〈x4,xD〉
+  | xC ⇒ 〈x5,x4〉   | xD ⇒ 〈x5,xB〉   | xE ⇒ 〈x6,x2〉   | xF ⇒ 〈x6,x9〉 ]
+ | x8 ⇒ match e2 with
+  [ x0 ⇒ 〈x0,x0〉   | x1 ⇒ 〈x0,x8〉   | x2 ⇒ 〈x1,x0〉   | x3 ⇒ 〈x1,x8〉
+  | x4 ⇒ 〈x2,x0〉   | x5 ⇒ 〈x2,x8〉   | x6 ⇒ 〈x3,x0〉   | x7 ⇒ 〈x3,x8〉
+  | x8 ⇒ 〈x4,x0〉   | x9 ⇒ 〈x4,x8〉   | xA ⇒ 〈x5,x0〉   | xB ⇒ 〈x5,x8〉
+  | xC ⇒ 〈x6,x0〉   | xD ⇒ 〈x6,x8〉   | xE ⇒ 〈x7,x0〉   | xF ⇒ 〈x7,x8〉 ]
+ | x9 ⇒ match e2 with
+  [ x0 ⇒ 〈x0,x0〉   | x1 ⇒ 〈x0,x9〉   | x2 ⇒ 〈x1,x2〉   | x3 ⇒ 〈x1,xB〉
+  | x4 ⇒ 〈x2,x4〉   | x5 ⇒ 〈x2,xD〉   | x6 ⇒ 〈x3,x6〉   | x7 ⇒ 〈x3,xF〉
+  | x8 ⇒ 〈x4,x8〉   | x9 ⇒ 〈x5,x1〉   | xA ⇒ 〈x5,xA〉   | xB ⇒ 〈x6,x3〉
+  | xC ⇒ 〈x6,xC〉   | xD ⇒ 〈x7,x5〉   | xE ⇒ 〈x7,xE〉   | xF ⇒ 〈x8,x7〉 ]
+ | xA ⇒ match e2 with
+  [ x0 ⇒ 〈x0,x0〉   | x1 ⇒ 〈x0,xA〉   | x2 ⇒ 〈x1,x4〉   | x3 ⇒ 〈x1,xE〉
+  | x4 ⇒ 〈x2,x8〉   | x5 ⇒ 〈x3,x2〉   | x6 ⇒ 〈x3,xC〉   | x7 ⇒ 〈x4,x6〉
+  | x8 ⇒ 〈x5,x0〉   | x9 ⇒ 〈x5,xA〉   | xA ⇒ 〈x6,x4〉   | xB ⇒ 〈x6,xE〉
+  | xC ⇒ 〈x7,x8〉   | xD ⇒ 〈x8,x2〉   | xE ⇒ 〈x8,xC〉   | xF ⇒ 〈x9,x6〉 ]
+ | xB ⇒ match e2 with
+  [ x0 ⇒ 〈x0,x0〉   | x1 ⇒ 〈x0,xB〉   | x2 ⇒ 〈x1,x6〉   | x3 ⇒ 〈x2,x1〉
+  | x4 ⇒ 〈x2,xC〉   | x5 ⇒ 〈x3,x7〉   | x6 ⇒ 〈x4,x2〉   | x7 ⇒ 〈x4,xD〉
+  | x8 ⇒ 〈x5,x8〉   | x9 ⇒ 〈x6,x3〉   | xA ⇒ 〈x6,xE〉   | xB ⇒ 〈x7,x9〉
+  | xC ⇒ 〈x8,x4〉   | xD ⇒ 〈x8,xF〉   | xE ⇒ 〈x9,xA〉   | xF ⇒ 〈xA,x5〉 ]
+ | xC ⇒ match e2 with
+  [ x0 ⇒ 〈x0,x0〉   | x1 ⇒ 〈x0,xC〉   | x2 ⇒ 〈x1,x8〉   | x3 ⇒ 〈x2,x4〉
+  | x4 ⇒ 〈x3,x0〉   | x5 ⇒ 〈x3,xC〉   | x6 ⇒ 〈x4,x8〉   | x7 ⇒ 〈x5,x4〉
+  | x8 ⇒ 〈x6,x0〉   | x9 ⇒ 〈x6,xC〉   | xA ⇒ 〈x7,x8〉   | xB ⇒ 〈x8,x4〉
+  | xC ⇒ 〈x9,x0〉   | xD ⇒ 〈x9,xC〉   | xE ⇒ 〈xA,x8〉   | xF ⇒ 〈xB,x4〉 ]
+ | xD ⇒ match e2 with
+  [ x0 ⇒ 〈x0,x0〉   | x1 ⇒ 〈x0,xD〉   | x2 ⇒ 〈x1,xA〉   | x3 ⇒ 〈x2,x7〉
+  | x4 ⇒ 〈x3,x4〉   | x5 ⇒ 〈x4,x1〉   | x6 ⇒ 〈x4,xE〉   | x7 ⇒ 〈x5,xB〉
+  | x8 ⇒ 〈x6,x8〉   | x9 ⇒ 〈x7,x5〉   | xA ⇒ 〈x8,x2〉   | xB ⇒ 〈x8,xF〉
+  | xC ⇒ 〈x9,xC〉   | xD ⇒ 〈xA,x9〉   | xE ⇒ 〈xB,x6〉   | xF ⇒ 〈xC,x3〉 ]
+ | xE ⇒ match e2 with
+  [ x0 ⇒ 〈x0,x0〉   | x1 ⇒ 〈x0,xE〉   | x2 ⇒ 〈x1,xC〉   | x3 ⇒ 〈x2,xA〉
+  | x4 ⇒ 〈x3,x8〉   | x5 ⇒ 〈x4,x6〉   | x6 ⇒ 〈x5,x4〉   | x7 ⇒ 〈x6,x2〉
+  | x8 ⇒ 〈x7,x0〉   | x9 ⇒ 〈x7,xE〉   | xA ⇒ 〈x8,xC〉   | xB ⇒ 〈x9,xA〉
+  | xC ⇒ 〈xA,x8〉   | xD ⇒ 〈xB,x6〉   | xE ⇒ 〈xC,x4〉   | xF ⇒ 〈xD,x2〉 ]
+ | xF ⇒ match e2 with
+  [ x0 ⇒ 〈x0,x0〉   | x1 ⇒ 〈x0,xF〉   | x2 ⇒ 〈x1,xE〉   | x3 ⇒ 〈x2,xD〉
+  | x4 ⇒ 〈x3,xC〉   | x5 ⇒ 〈x4,xB〉   | x6 ⇒ 〈x5,xA〉   | x7 ⇒ 〈x6,x9〉
+  | x8 ⇒ 〈x7,x8〉   | x9 ⇒ 〈x8,x7〉   | xA ⇒ 〈x9,x6〉   | xB ⇒ 〈xA,x5〉
+  | xC ⇒ 〈xB,x4〉   | xD ⇒ 〈xC,x3〉   | xE ⇒ 〈xD,x2〉   | xF ⇒ 〈xE,x1〉 ]
+ ].
+
+(* correzione per somma su BCD *)
+(* input: halfcarry,carry,X(BCD+BCD) *)
+(* output: X',carry' *)
+definition daa_b8 ≝
+λh,c:bool.λX:byte8.
+ match lt_b8 X 〈x9,xA〉 with
+  (* [X:0x00-0x99] *)
+  (* c' = c *)
+  (* X' = [(b16l X):0x0-0x9] X + [h=1 ? 0x06 : 0x00] + [c=1 ? 0x60 : 0x00]
+          [(b16l X):0xA-0xF] X + 0x06 + [c=1 ? 0x60 : 0x00] *)
+  [ true ⇒
+   let X' ≝ match (lt_ex (b8l X) xA) ⊗ (⊖h) with
+    [ true ⇒ X
+    | false ⇒ plus_b8nc X 〈x0,x6〉 ] in
+   let X'' ≝ match c with
+    [ true ⇒ plus_b8nc X' 〈x6,x0〉
+    | false ⇒ X' ] in
+   pairT ?? X'' c
+  (* [X:0x9A-0xFF] *)
+  (* c' = 1 *)
+  (* X' = [X:0x9A-0xFF]
+          [(b16l X):0x0-0x9] X + [h=1 ? 0x06 : 0x00] + 0x60
+          [(b16l X):0xA-0xF] X + 0x6 + 0x60 *) 
+  | false ⇒
+   let X' ≝ match (lt_ex (b8l X) xA) ⊗ (⊖h) with
+    [ true ⇒ X
+    | false ⇒ plus_b8nc X 〈x0,x6〉 ] in
+   let X'' ≝ plus_b8nc X' 〈x6,x0〉 in
+   pairT ?? X'' true
+  ].
+
+(* iteratore sui byte *)
+definition forall_byte8 ≝
+ λP.
+  forall_exadecim (λbh.
+  forall_exadecim (λbl.
+   P (mk_byte8 bh bl))).
+
+(* ********************** *)
+(* TEOREMI/LEMMMI/ASSIOMI *)
+(* ********************** *)
+
+lemma byte8_of_nat_nat_of_byte8: ∀b. byte8_of_nat (nat_of_byte8 b) = b.
+ intros;
+ elim b;
+ elim e;
+ elim e1;
+ reflexivity.
+qed.
+
+lemma lt_nat_of_byte8_256: ∀b. nat_of_byte8 b < 256.
+ intro;
+ unfold nat_of_byte8;
+ letin H ≝ (lt_nat_of_exadecim_16 (b8h b)); clearbody H;
+ letin K ≝ (lt_nat_of_exadecim_16 (b8l b)); clearbody K;
+ unfold lt in H K ⊢ %;
+ letin H' ≝ (le_S_S_to_le ? ? H); clearbody H'; clear H;
+ letin K' ≝ (le_S_S_to_le ? ? K); clearbody K'; clear K;
+ apply le_S_S;
+ cut (16*b8h b ≤ 16*15);
+  [ letin Hf ≝ (le_plus ? ? ? ? Hcut K'); clearbody Hf;
+    simplify in Hf:(? ? %);
+    assumption
+  | apply le_times_r. apply H'.
+  ]
+qed.
+
+lemma nat_of_byte8_byte8_of_nat: ∀n. nat_of_byte8 (byte8_of_nat n) = n \mod 256.
+ intro;
+ letin H ≝ (lt_nat_of_byte8_256 (byte8_of_nat n)); clearbody H;
+ rewrite < (lt_to_eq_mod ? ? H); clear H;
+ unfold byte8_of_nat;
+ unfold nat_of_byte8;
+ change with ((16*(exadecim_of_nat (n/16)) + exadecim_of_nat n) \mod 256 = n \mod 256);
+ letin H ≝ (div_mod n 16 ?); clearbody H; [ autobatch | ];
+ rewrite > symmetric_times in H;
+ rewrite > nat_of_exadecim_exadecim_of_nat in ⊢ (? ? (? (? % ?) ?) ?);
+ rewrite > nat_of_exadecim_exadecim_of_nat in ⊢ (? ? (? (? ? %) ?) ?);
+ rewrite > H in ⊢ (? ? ? (? % ?)); clear H;
+ rewrite > mod_plus in ⊢ (? ? % ?);
+ rewrite > mod_plus in ⊢ (? ? ? %);
+ apply eq_mod_to_eq_plus_mod;
+ rewrite < mod_mod in ⊢ (? ? ? %); [ | autobatch];
+ rewrite < mod_mod in ⊢ (? ? % ?); [ | autobatch];
+ rewrite < (eq_mod_times_times_mod ? ? 16 256) in ⊢ (? ? (? % ?) ?); [2: reflexivity | ];
+ rewrite < mod_mod in ⊢ (? ? % ?);
+  [ reflexivity
+  | autobatch
+  ].
+qed.
+
+lemma eq_nat_of_byte8_n_nat_of_byte8_mod_n_256:
+ ∀n. byte8_of_nat n = byte8_of_nat (n \mod 256).
+ intro;
+ unfold byte8_of_nat;
+ apply eq_f2;
+  [ rewrite > exadecim_of_nat_mod in ⊢ (? ? % ?);
+    rewrite > exadecim_of_nat_mod in ⊢ (? ? ? %);
+    apply eq_f;
+    elim daemon
+  | rewrite > exadecim_of_nat_mod;
+    rewrite > exadecim_of_nat_mod in ⊢ (? ? ? %);
+    rewrite > divides_to_eq_mod_mod_mod;
+     [ reflexivity
+     | apply (witness ? ? 16). reflexivity.
+     ]
+  ]
+qed.
+
+lemma plusb8_ok:
+ ∀b1,b2,c.
+  match plus_b8 b1 b2 c with
+   [ pairT r c' ⇒ b1 + b2 + nat_of_bool c = nat_of_byte8 r + nat_of_bool c' * 256
+   ].
+ intros; elim daemon.
+qed.
+
+lemma plusb8_O_x:
+ ∀b. plus_b8 (mk_byte8 x0 x0) b false = pairT ?? b false.
+ intros;
+ elim b;
+ elim e;
+ elim e1;
+ reflexivity.
+qed.
+
+lemma plusb8nc_O_x:
+ ∀x. plus_b8nc (mk_byte8 x0 x0) x = x.
+ intros;
+ unfold plus_b8nc;
+ rewrite > plusb8_O_x;
+ reflexivity.
+qed.
+
+lemma eq_nat_of_byte8_mod: ∀b. nat_of_byte8 b = nat_of_byte8 b \mod 256.
+ intro;
+ lapply (lt_nat_of_byte8_256 b);
+ rewrite > (lt_to_eq_mod ? ? Hletin) in ⊢ (? ? ? %);
+ reflexivity.
+qed.
+
+theorem plusb8nc_ok:
+ ∀b1,b2:byte8.nat_of_byte8 (plus_b8nc b1 b2) = (b1 + b2) \mod 256.
+ intros;
+ unfold plus_b8nc;
+ generalize in match (plusb8_ok b1 b2 false);
+ elim (plus_b8 b1 b2 false);
+ simplify in H ⊢ %;
+ change with (nat_of_byte8 t = (b1 + b2) \mod 256);
+ rewrite < plus_n_O in H;
+ rewrite > H; clear H;
+ rewrite > mod_plus;
+ letin K ≝ (eq_nat_of_byte8_mod t); clearbody K;
+ letin K' ≝ (eq_mod_times_n_m_m_O (nat_of_bool t1) 256 ?); clearbody K';
+  [ autobatch | ];
+ autobatch paramodulation.
+qed.
+
+lemma eq_eqb8_x0_x0_byte8_of_nat_S_false:
+ ∀b. b < 255 → eq_b8 (mk_byte8 x0 x0) (byte8_of_nat (S b)) = false.
+ intros;
+ unfold byte8_of_nat;
+ cut (b < 15 ∨ b ≥ 15);
+  [ elim Hcut;
+    [ unfold eq_b8;
+      change in ⊢ (? ? (? ? %) ?) with (eq_ex x0 (exadecim_of_nat (S b))); 
+      rewrite > eq_eqex_S_x0_false;
+       [ elim (eq_ex (b8h (mk_byte8 x0 x0))
+          (b8h (mk_byte8 (exadecim_of_nat (S b/16)) (exadecim_of_nat (S b)))));
+         simplify;
+         reflexivity
+       | assumption
+       ]
+    | unfold eq_b8;
+      change in ⊢ (? ? (? % ?) ?) with (eq_ex x0 (exadecim_of_nat (S b/16)));
+      letin K ≝ (leq_m_n_to_eq_div_n_m_S (S b) 16 ? ?);
+       [ autobatch
+       | unfold in H1;
+         apply le_S_S;
+         assumption
+       | clearbody K;
+         elim K; clear K;
+         rewrite > H2;
+         rewrite > eq_eqex_S_x0_false;
+          [ reflexivity
+          | unfold lt;
+            unfold lt in H;
+            rewrite < H2;
+            clear H2; clear a; clear H1; clear Hcut;
+            apply (le_times_to_le 16) [ autobatch | ] ;
+            rewrite > (div_mod (S b) 16) in H;[2:autobatch|]
+            rewrite > (div_mod 255 16) in H:(? ? %);[2:autobatch|]
+            lapply (le_to_le_plus_to_le ? ? ? ? ? H);
+            [apply lt_S_to_le;
+             apply lt_mod_m_m;autobatch
+            |rewrite > sym_times;
+             rewrite > sym_times in ⊢ (? ? %);
+             normalize in ⊢ (? ? %);apply Hletin;
+            ]
+          ] 
+       ]
+    ]
+  | elim (or_lt_le b 15);unfold ge;autobatch
+  ].
+qed.
+
+axiom eq_mod_O_to_exists: ∀n,m. n \mod m = 0 → ∃z. n = z*m.
+
+lemma eq_b8pred_S_a_a:
+ ∀a. a < 255 → pred_b8 (byte8_of_nat (S a)) = byte8_of_nat a.
+ intros;
+ unfold pred_b8;
+ apply (bool_elim ? (eq_ex (b8l (byte8_of_nat (S a))) x0)); intros;
+  [ change with (mk_byte8 (pred_ex (b8h (byte8_of_nat (S a)))) (pred_ex (b8l (byte8_of_nat (S a))))
+     = byte8_of_nat a);
+    rewrite > (eqex_true_to_eq ? ? H1);
+    normalize in ⊢ (? ? (? ? %) ?);
+    unfold byte8_of_nat;
+    change with (mk_byte8 (pred_ex (exadecim_of_nat (S a/16))) xF =
+                 mk_byte8 (exadecim_of_nat (a/16)) (exadecim_of_nat a));
+    lapply (eqex_true_to_eq ? ? H1); clear H1;
+    unfold byte8_of_nat in Hletin;
+    change in Hletin with (exadecim_of_nat (S a) = x0);
+    lapply (eq_f ? ? nat_of_exadecim ? ? Hletin); clear Hletin;
+    normalize in Hletin1:(? ? ? %);
+    rewrite > nat_of_exadecim_exadecim_of_nat in Hletin1;
+    elim (eq_mod_O_to_exists ? ? Hletin1); clear Hletin1;
+    rewrite > H1;
+    rewrite > lt_O_to_div_times; [2: autobatch | ]
+    lapply (eq_f ? ? (λx.x/16) ? ? H1);
+    rewrite > lt_O_to_div_times in Hletin; [2: autobatch | ]
+    lapply (eq_f ? ? (λx.x \mod 16) ? ? H1);
+    rewrite > eq_mod_times_n_m_m_O in Hletin1;
+    elim daemon
+  | change with (mk_byte8 (b8h (byte8_of_nat (S a))) (pred_ex (b8l (byte8_of_nat (S a))))
+    = byte8_of_nat a);
+    unfold byte8_of_nat;
+    change with (mk_byte8 (exadecim_of_nat (S a/16)) (pred_ex (exadecim_of_nat (S a)))
+    = mk_byte8 (exadecim_of_nat (a/16)) (exadecim_of_nat a));
+    lapply (eqex_false_to_not_eq ? ? H1);
+    unfold byte8_of_nat in Hletin;
+    change in Hletin with (exadecim_of_nat (S a) ≠ x0);
+    cut (nat_of_exadecim (exadecim_of_nat (S a)) ≠ 0);
+     [2: intro;
+       apply Hletin;
+       lapply (eq_f ? ? exadecim_of_nat ? ? H2);
+       rewrite > exadecim_of_nat_nat_of_exadecim in Hletin1;
+       apply Hletin1
+     | ];
+    elim daemon
+  ]
+qed.
+
+lemma plusb8nc_S:
+ ∀x:byte8.∀n.plus_b8nc (byte8_of_nat (x*n)) x = byte8_of_nat (x * S n).
+ intros;
+ rewrite < byte8_of_nat_nat_of_byte8;
+ rewrite > (plusb8nc_ok (byte8_of_nat (x*n)) x);
+ rewrite < times_n_Sm;
+ rewrite > nat_of_byte8_byte8_of_nat in ⊢ (? ? (? (? (? % ?) ?)) ?);
+ rewrite > eq_nat_of_byte8_n_nat_of_byte8_mod_n_256 in ⊢ (? ? ? %);
+ rewrite > mod_plus in ⊢ (? ? (? %) ?);
+ rewrite > mod_plus in ⊢ (? ? ? (? %));
+ rewrite < mod_mod in ⊢ (? ? (? (? (? % ?) ?)) ?); [2: autobatch | ];
+ rewrite > sym_plus in ⊢ (? ? (? (? % ?)) ?);
+ reflexivity.
+qed.
+
+lemma eq_plusb8c_x0_x0_x_false:
+ ∀x.plus_b8c (mk_byte8 x0 x0) x = false.
+ intro;
+ elim x;
+ elim e;
+ elim e1;
+ reflexivity.
+qed.
+
+axiom eqb8_true_to_eq: ∀b,b'. eq_b8 b b' = true → b=b'.
+
+axiom eqb8_false_to_not_eq: ∀b,b'. eq_b8 b b' = false → b ≠ b'.
+
+axiom byte8_of_nat_mod: ∀n.byte8_of_nat n = byte8_of_nat (n \mod 256).
+
+(* nuovi *)
+
+lemma ok_mul_ex :
+∀e1,e2.nat_of_byte8 (mul_ex e1 e2) = (nat_of_exadecim e1) * (nat_of_exadecim e2).
+intros;
+elim e1;
+elim e2;
+reflexivity.
+qed.
diff --git a/helm/software/matita/library/freescale/doc/aurei.txt b/helm/software/matita/library/freescale/doc/aurei.txt
new file mode 100644 (file)
index 0000000..dc62cf6
--- /dev/null
@@ -0,0 +1,235 @@
+\r
+static unsigned int result[16]={ 0,0,0,0,0,0,0,0,0,0,0,0,0,0,0,0 };\r
+\r
+void main(void)\r
+{\r
+unsigned int res_pos=0,tested_num=0,divisor=0;\r
+unsigned long int acc=0;\r
+\r
+for(tested_num=1;tested_num<2;tested_num++)\r
+ {\r
+ for(acc=0,divisor=1;divisor<tested_num;divisor++)\r
+  {\r
+  if(!(tested_num%divisor))\r
+   { acc+=divisor; }\r
+  }\r
\r
+ if(acc==tested_num)\r
+  { result[res_pos++]=tested_num; }\r
+ }\r
+}\r
+\r
+NUMERI AUREI: Somma divisori(x)=x\r
+Fino a 0xFFFF sono\r
+6/28/496/8128\r
+\r
+ 1000: 545 -> 156.295.006          : stima di complessita' che dovrebbe essere\r
+  500: 545 ->  27.863.083          : T(n)=O(a+b*(n^3)) ma e' veramente enorme!\r
+  250: 545 ->   4.916.615          :\r
+  100: 545 ->     826.150          :\r
+   50: 545 ->     218.243 (213878) : 217698=64+217634 : /117649= 1,8\r
+   25: 545 ->      58.798 ( 57858) :  58253=64+ 58189 : / 13824= 4,2\r
+   10: 545 ->      10.212 ( 10127) :   9667=64+  9603 : /   729=13,2\r
+    5: 545 ->       2.789 (  2789) :   2244=64+  2180 : /    64=34,1\r
+    4: 545 ->       1.760 (  1765) :   1220=64+  1156 : /    27=42,8\r
+    3: 545 ->       1.136 (  1142) :    597=64+   533 : /    8 =66,6\r
+    2: 545 ->         661 (   664) :    119=64+    55 : /    1 =55\r
+    1: 545 ->         609 (   601) :        64        :\r
+\r
+word result[16] = 0x0100\r
+\r
+A: 0x00 HX: 0x1A00 PC: 0x18C0 SP: 0x0165 Z:1\r
+A: 0x00 HX: num    PC: 0x1941 SP: 0x0165 Z:1\r
+\r
+18C0 95          TSX\r
+18C1 6F 09       CLR 9,X\r
+18C3 6F 08       CLR 8,X\r
+18C5 AD 7D       BSR *+127 ; 1944\r
+18C7 9E FF 01    STHX 1,SP\r
+18CA AD 78       BSR *+122 ; 1944\r
+18CC 9E FF 07    STHX 7,SP\r
+18CF 20 38       BRA *+58\r
+18D1 9E FE 01    LDHX 1,SP\r
+18D4 89          PSHX\r
+18D5 8B          PSHH\r
+18D6 9E FE 09    LDHX 9,SP\r
+18D9 CD 1A 0D    JSR 0X1A0D\r
+18DC A7 02       AIS #2\r
+18DE 65 00 00    CPHX #0X0000\r
+18E1 26 1F       BNE *+33 ; 1902\r
+18E3 95          TSX\r
+18E4 E6 07       LDA 7,X\r
+18E6 EE 06       LDX 6,X\r
+18E8 87          PSHA\r
+18E9 89          PSHX\r
+18EA 4F          CLRA\r
+18EB 87          PSHA\r
+18EC 87          PSHA\r
+18ED 95          TSX\r
+18EE 89          PSHX\r
+18EF 8B          PSHH\r
+18F0 AF 06       AIX #6\r
+18F2 89          PSHX\r
+18F3 8B          PSHH\r
+18F4 9E FE 03    LDHX 3,SP\r
+18F7 CD 1A 1D    JSR 0X1A1D\r
+18FA 95          TSX\r
+18FB AF 0C       AIX #12\r
+18FD CD 1A 23    JSR 0X1A23\r
+1900 A7 06       AIS #6\r
+1902 95          TSX\r
+1903 6C 07       INC 7,X\r
+1905 26 02       BNE *+4 ; 1909\r
+1907 6C 06       INC 6,X\r
+1909 9E FE 01    LDHX 1,SP\r
+190C 9E F3 07    CPHX 7,SP\r
+190F 22 C0       BHI *-62 ; 18D1\r
+1911 9E F3 05    CPHX 5,SP\r
+1914 26 1D       BNE *+31 ; 1933\r
+1916 9E FE 03    LDHX 3,SP\r
+1919 26 18       BNE *+26 ; 1933\r
+191B 9E FE 09    LDHX 9,SP\r
+191E 89          PSHX\r
+191F AF 01       AIX #1\r
+1921 9E FF 0A    STHX 10,SP\r
+1924 88          PULX\r
+1925 58          LSLX\r
+1926 9E E6 02    LDA 2,SP\r
+1929 8C          CLRH\r
+192A D7 01 01    STA 257,X\r
+192D 9E E6 01    LDA 1,SP\r
+1930 D7 01 00    STA 256,X\r
+1933 95          TSX\r
+1934 6C 01       INC 1,X\r
+1936 26 01       BNE *+3 ; 1939\r
+1938 7C          INC ,X\r
+1939 9E FE 01    LDHX 1,SP\r
+193C 65 00 19    CPHX #0X0019\r
+193F 25 89       BCS *-117 ; 18CA\r
+\r
+1941-1943 STOP\r
+\r
+1944 5F          CLRX\r
+1945 8C          CLRH\r
+1946 9E FF 07    STHX 7,SP\r
+1949 8C          CLRH\r
+194A 9E FF 05    STHX 5,SP\r
+194D 5C          INCX\r
+194E 81          RTS\r
+194F E6 03       LDA 3,X\r
+1951 87          PSHA\r
+1952 E6 02       LDA 2,X\r
+1954 87          PSHA\r
+1955 9E AE       LDHX ,X\r
+1957 89          PSHX\r
+1958 8B          PSHH\r
+1959 9E FE 07    LDHX 7,SP\r
+195C E6 03       LDA 3,X\r
+195E 9E E7 11    STA 17,SP\r
+1961 E6 02       LDA 2,X\r
+1963 9E E7 10    STA 16,SP\r
+1966 9E AE       LDHX ,X\r
+1968 9E FF 0E    STHX 14,SP\r
+196B 9E FE 05    LDHX 5,SP\r
+196E FC          JMP ,X\r
+196F 87          PSHA\r
+1970 89          PSHX\r
+1971 8B          PSHH\r
+1972 89          PSHX\r
+1973 8B          PSHH\r
+1974 9E FE 06    LDHX 6,SP\r
+1977 89          PSHX\r
+1978 8B          PSHH\r
+1979 9E FE 0A    LDHX 10,SP\r
+197C 9E FF 08    STHX 8,SP\r
+197F 9E FE 0C    LDHX 12,SP\r
+1982 CC 19 4F    JMP 0X194F\r
+1985 9E 6D 04    TST 4,SP\r
+1988 26 1A       BNE *+28 ; 19A4\r
+198A 95          TSX\r
+198B E6 07       LDA 7,X\r
+198D EE 04       LDX 4,X\r
+198F 8C          CLRH\r
+1990 52          DIV\r
+1991 9E E7 04    STA 4,SP\r
+1994 9E E6 09    LDA 9,SP\r
+1997 52          DIV\r
+1998 9E E7 05    STA 5,SP\r
+199B 9E 6F 08    CLR 8,SP\r
+199E 8B          PSHH\r
+199F 86          PULA\r
+19A0 9E E7 09    STA 9,SP\r
+19A3 81          RTS\r
+19A4 4F          CLRA\r
+19A5 87          PSHA\r
+19A6 AE 08       LDX #0X08\r
+19A8 98          CLC\r
+19A9 9E 69 0A    ROL 10,SP\r
+19AC 9E 69 09    ROL 9,SP\r
+19AF 9E 69 01    ROL 1,SP\r
+19B2 9E E6 05    LDA 5,SP\r
+19B5 9E E1 01    CMP 1,SP\r
+19B8 22 1D       BHI *+31 ; 19D7\r
+19BA 26 08       BNE *+10 ; 19C4\r
+19BC 9E E6 06    LDA 6,SP\r
+19BF 9E E1 09    CMP 9,SP\r
+19C2 22 13       BHI *+21 ; 19D7\r
+19C4 9E E6 09    LDA 9,SP\r
+19C7 9E E0 06    SUB 6,SP\r
+19CA 9E E7 09    STA 9,SP\r
+19CD 9E E6 01    LDA 1,SP\r
+19D0 9E E2 05    SBC 5,SP\r
+19D3 9E E7 01    STA 1,SP\r
+19D6 99          SEC\r
+19D7 5B D0       DBNZX *-46 ; 19A9\r
+19D9 9E E6 0A    LDA 10,SP\r
+19DC 49          ROLA\r
+19DD 9E E7 06    STA 6,SP\r
+19E0 9E E6 09    LDA 9,SP\r
+19E3 9E E7 0A    STA 10,SP\r
+19E6 86          PULA\r
+19E7 9E E7 08    STA 8,SP\r
+19EA 9E 6F 04    CLR 4,SP\r
+19ED 81          RTS\r
+19EE 95          TSX\r
+19EF E6 12       LDA 18,X\r
+19F1 EB 05       ADD 5,X\r
+19F3 E7 12       STA 18,X\r
+19F5 E6 11       LDA 17,X\r
+19F7 E9 04       ADC 4,X\r
+19F9 E7 11       STA 17,X\r
+19FB E6 10       LDA 16,X\r
+19FD E9 03       ADC 3,X\r
+19FF E7 10       STA 16,X\r
+1A01 E6 0F       LDA 15,X\r
+1A03 E9 02       ADC 2,X\r
+1A05 E7 0F       STA 15,X\r
+1A07 A7 0A       AIS #10\r
+1A09 8A          PULH\r
+1A0A 88          PULX\r
+1A0B 86          PULA\r
+1A0C 81          RTS\r
+1A0D A7 FE       AIS #-2\r
+1A0F 9E FF 01    STHX 1,SP\r
+1A12 87          PSHA\r
+1A13 CD 19 85    JSR 0X1985\r
+1A16 86          PULA\r
+1A17 A7 02       AIS #2\r
+1A19 9E FE 03    LDHX 3,SP\r
+1A1C 81          RTS\r
+1A1D CD 19 6F    JSR 0X196F\r
+1A20 CD 19 EE    JSR 0X19EE\r
+1A23 87          PSHA\r
+1A24 9E E6 04    LDA 4,SP\r
+1A27 F7          STA ,X\r
+1A28 9E E6 05    LDA 5,SP\r
+1A2B E7 01       STA 1,X\r
+1A2D 9E E6 06    LDA 6,SP\r
+1A30 E7 02       STA 2,X\r
+1A32 9E E6 07    LDA 7,SP\r
+1A35 E7 03       STA 3,X\r
+1A37 86          PULA\r
+1A38 8A          PULH\r
+1A39 88          PULX\r
+1A3A A7 A4       AIS #4\r
+1A3C FC          JMP ,X\r
diff --git a/helm/software/matita/library/freescale/doc/daa.txt b/helm/software/matita/library/freescale/doc/daa.txt
new file mode 100644 (file)
index 0000000..9babf07
--- /dev/null
@@ -0,0 +1,1334 @@
+sorgente utilizzato per produrre la tabella:\r
+\r
+unsigned int i;\r
+unsigned char val;\r
+\r
+for(i=0,val=0;i<0x100;i++,val++)\r
+ {\r
+ asm {\r
+  /* H=0 C=0 */\r
+  lda val\r
+  psha\r
+  tap\r
+  and #0x00\r
+  tpa\r
+  pula\r
+  daa\r
+  nop <- breakpoint/prelevare A,SR\r
+\r
+  /* H=0 C=1 */\r
+  lda val\r
+  psha\r
+  tap\r
+  and #0x00\r
+  or  #0x01\r
+  tpa\r
+  pula\r
+  daa\r
+  nop <- breakpoint/prelevare A,SR\r
+\r
+  /* H=1 C=0 */\r
+  lda val\r
+  psha\r
+  tap\r
+  and #0x00\r
+  or  #0x10\r
+  tpa\r
+  pula\r
+  daa\r
+  nop <- breakpoint/prelevare A,SR\r
+\r
+  /* H=1 C=1 */\r
+  lda val\r
+  psha\r
+  tap\r
+  and #0x00\r
+  or  #0x11\r
+  tpa\r
+  pula\r
+  daa\r
+  nop <- breakpoint/prelevare A,SR\r
+  }\r
+ }\r
+\r
+************************************************\r
+\r
+H(0) C(0) A(0x00) -> V(0) N(0) Z(1) C(0) A(0x00)\r
+H(0) C(1) A(0x00) -> V(0) N(0) Z(0) C(1) A(0x60)\r
+H(1) C(0) A(0x00) -> V(0) N(0) Z(0) C(0) A(0x06)\r
+H(1) C(1) A(0x00) -> V(0) N(0) Z(0) C(1) A(0x66)\r
+\r
+H(0) C(0) A(0x01) -> V(0) N(0) Z(0) C(0) A(0x01)\r
+H(0) C(1) A(0x01) -> V(0) N(0) Z(0) C(1) A(0x61)\r
+H(1) C(0) A(0x01) -> V(0) N(0) Z(0) C(0) A(0x07)\r
+H(1) C(1) A(0x01) -> V(0) N(0) Z(0) C(1) A(0x67)\r
+\r
+H(0) C(0) A(0x02) -> V(0) N(0) Z(0) C(0) A(0x02)\r
+H(0) C(1) A(0x02) -> V(0) N(0) Z(0) C(1) A(0x62)\r
+H(1) C(0) A(0x02) -> V(0) N(0) Z(0) C(0) A(0x08)\r
+H(1) C(1) A(0x02) -> V(0) N(0) Z(0) C(1) A(0x68)\r
+\r
+H(0) C(0) A(0x03) -> V(0) N(0) Z(0) C(0) A(0x03)\r
+H(0) C(1) A(0x03) -> V(0) N(0) Z(0) C(1) A(0x63)\r
+H(1) C(0) A(0x03) -> V(0) N(0) Z(0) C(0) A(0x09)\r
+H(1) C(1) A(0x03) -> V(0) N(0) Z(0) C(1) A(0x69)\r
+\r
+H(0) C(0) A(0x04) -> V(0) N(0) Z(0) C(0) A(0x04)\r
+H(0) C(1) A(0x04) -> V(0) N(0) Z(0) C(1) A(0x64)\r
+H(1) C(0) A(0x04) -> V(0) N(0) Z(0) C(0) A(0x0A)\r
+H(1) C(1) A(0x04) -> V(0) N(0) Z(0) C(1) A(0x6A)\r
+\r
+H(0) C(0) A(0x05) -> V(0) N(0) Z(0) C(0) A(0x05)\r
+H(0) C(1) A(0x05) -> V(0) N(0) Z(0) C(1) A(0x65)\r
+H(1) C(0) A(0x05) -> V(0) N(0) Z(0) C(0) A(0x0B)\r
+H(1) C(1) A(0x05) -> V(0) N(0) Z(0) C(1) A(0x6B)\r
+\r
+H(0) C(0) A(0x06) -> V(0) N(0) Z(0) C(0) A(0x06)\r
+H(0) C(1) A(0x06) -> V(0) N(0) Z(0) C(1) A(0x66)\r
+H(1) C(0) A(0x06) -> V(0) N(0) Z(0) C(0) A(0x0C)\r
+H(1) C(1) A(0x06) -> V(0) N(0) Z(0) C(1) A(0x6C)\r
+\r
+H(0) C(0) A(0x07) -> V(0) N(0) Z(0) C(0) A(0x07)\r
+H(0) C(1) A(0x07) -> V(0) N(0) Z(0) C(1) A(0x67)\r
+H(1) C(0) A(0x07) -> V(0) N(0) Z(0) C(0) A(0x0D)\r
+H(1) C(1) A(0x07) -> V(0) N(0) Z(0) C(1) A(0x6D)\r
+\r
+H(0) C(0) A(0x08) -> V(0) N(0) Z(0) C(0) A(0x08)\r
+H(0) C(1) A(0x08) -> V(0) N(0) Z(0) C(1) A(0x68)\r
+H(1) C(0) A(0x08) -> V(0) N(0) Z(0) C(0) A(0x0E)\r
+H(1) C(1) A(0x08) -> V(0) N(0) Z(0) C(1) A(0x6E)\r
+\r
+H(0) C(0) A(0x09) -> V(0) N(0) Z(0) C(0) A(0x09)\r
+H(0) C(1) A(0x09) -> V(0) N(0) Z(0) C(1) A(0x69)\r
+H(1) C(0) A(0x09) -> V(0) N(0) Z(0) C(0) A(0x0F)\r
+H(1) C(1) A(0x09) -> V(0) N(0) Z(0) C(1) A(0x6F)\r
+\r
+H(0) C(0) A(0x0A) -> V(0) N(0) Z(0) C(0) A(0x10)\r
+H(0) C(1) A(0x0A) -> V(0) N(0) Z(0) C(1) A(0x70)\r
+H(1) C(0) A(0x0A) -> V(0) N(0) Z(0) C(0) A(0x10)\r
+H(1) C(1) A(0x0A) -> V(0) N(0) Z(0) C(1) A(0x70)\r
+\r
+H(0) C(0) A(0x0B) -> V(0) N(0) Z(0) C(0) A(0x11)\r
+H(0) C(1) A(0x0B) -> V(0) N(0) Z(0) C(1) A(0x71)\r
+H(1) C(0) A(0x0B) -> V(0) N(0) Z(0) C(0) A(0x11)\r
+H(1) C(1) A(0x0B) -> V(0) N(0) Z(0) C(1) A(0x71)\r
+\r
+H(0) C(0) A(0x0C) -> V(0) N(0) Z(0) C(0) A(0x12)\r
+H(0) C(1) A(0x0C) -> V(0) N(0) Z(0) C(1) A(0x72)\r
+H(1) C(0) A(0x0C) -> V(0) N(0) Z(0) C(0) A(0x12)\r
+H(1) C(1) A(0x0C) -> V(0) N(0) Z(0) C(1) A(0x72)\r
+\r
+H(0) C(0) A(0x0D) -> V(0) N(0) Z(0) C(0) A(0x13)\r
+H(0) C(1) A(0x0D) -> V(0) N(0) Z(0) C(1) A(0x73)\r
+H(1) C(0) A(0x0D) -> V(0) N(0) Z(0) C(0) A(0x13)\r
+H(1) C(1) A(0x0D) -> V(0) N(0) Z(0) C(1) A(0x73)\r
+\r
+H(0) C(0) A(0x0E) -> V(0) N(0) Z(0) C(0) A(0x14)\r
+H(0) C(1) A(0x0E) -> V(0) N(0) Z(0) C(1) A(0x74)\r
+H(1) C(0) A(0x0E) -> V(0) N(0) Z(0) C(0) A(0x14)\r
+H(1) C(1) A(0x0E) -> V(0) N(0) Z(0) C(1) A(0x74)\r
+\r
+H(0) C(0) A(0x0F) -> V(0) N(0) Z(0) C(0) A(0x15)\r
+H(0) C(1) A(0x0F) -> V(0) N(0) Z(0) C(1) A(0x75)\r
+H(1) C(0) A(0x0F) -> V(0) N(0) Z(0) C(0) A(0x15)\r
+H(1) C(1) A(0x0F) -> V(0) N(0) Z(0) C(1) A(0x75)\r
+\r
+H(0) C(0) A(0x10) -> V(0) N(0) Z(0) C(0) A(0x10)\r
+H(0) C(1) A(0x10) -> V(0) N(0) Z(0) C(1) A(0x70)\r
+H(1) C(0) A(0x10) -> V(0) N(0) Z(0) C(0) A(0x16)\r
+H(1) C(1) A(0x10) -> V(0) N(0) Z(0) C(1) A(0x76)\r
+\r
+H(0) C(0) A(0x11) -> V(0) N(0) Z(0) C(0) A(0x11)\r
+H(0) C(1) A(0x11) -> V(0) N(0) Z(0) C(1) A(0x71)\r
+H(1) C(0) A(0x11) -> V(0) N(0) Z(0) C(0) A(0x17)\r
+H(1) C(1) A(0x11) -> V(0) N(0) Z(0) C(1) A(0x77)\r
+\r
+H(0) C(0) A(0x12) -> V(0) N(0) Z(0) C(0) A(0x12)\r
+H(0) C(1) A(0x12) -> V(0) N(0) Z(0) C(1) A(0x72)\r
+H(1) C(0) A(0x12) -> V(0) N(0) Z(0) C(0) A(0x18)\r
+H(1) C(1) A(0x12) -> V(0) N(0) Z(0) C(1) A(0x78)\r
+\r
+H(0) C(0) A(0x13) -> V(0) N(0) Z(0) C(0) A(0x13)\r
+H(0) C(1) A(0x13) -> V(0) N(0) Z(0) C(1) A(0x73)\r
+H(1) C(0) A(0x13) -> V(0) N(0) Z(0) C(0) A(0x19)\r
+H(1) C(1) A(0x13) -> V(0) N(0) Z(0) C(1) A(0x79)\r
+\r
+H(0) C(0) A(0x14) -> V(0) N(0) Z(0) C(0) A(0x14)\r
+H(0) C(1) A(0x14) -> V(0) N(0) Z(0) C(1) A(0x74)\r
+H(1) C(0) A(0x14) -> V(0) N(0) Z(0) C(0) A(0x1A)\r
+H(1) C(1) A(0x14) -> V(0) N(0) Z(0) C(1) A(0x7A)\r
+\r
+H(0) C(0) A(0x15) -> V(0) N(0) Z(0) C(0) A(0x15)\r
+H(0) C(1) A(0x15) -> V(0) N(0) Z(0) C(1) A(0x75)\r
+H(1) C(0) A(0x15) -> V(0) N(0) Z(0) C(0) A(0x1B)\r
+H(1) C(1) A(0x15) -> V(0) N(0) Z(0) C(1) A(0x7B)\r
+\r
+H(0) C(0) A(0x16) -> V(0) N(0) Z(0) C(0) A(0x16)\r
+H(0) C(1) A(0x16) -> V(0) N(0) Z(0) C(1) A(0x76)\r
+H(1) C(0) A(0x16) -> V(0) N(0) Z(0) C(0) A(0x1C)\r
+H(1) C(1) A(0x16) -> V(0) N(0) Z(0) C(1) A(0x7C)\r
+\r
+H(0) C(0) A(0x17) -> V(0) N(0) Z(0) C(0) A(0x17)\r
+H(0) C(1) A(0x17) -> V(0) N(0) Z(0) C(1) A(0x77)\r
+H(1) C(0) A(0x17) -> V(0) N(0) Z(0) C(0) A(0x1D)\r
+H(1) C(1) A(0x17) -> V(0) N(0) Z(0) C(1) A(0x7D)\r
+\r
+H(0) C(0) A(0x18) -> V(0) N(0) Z(0) C(0) A(0x18)\r
+H(0) C(1) A(0x18) -> V(0) N(0) Z(0) C(1) A(0x78)\r
+H(1) C(0) A(0x18) -> V(0) N(0) Z(0) C(0) A(0x1E)\r
+H(1) C(1) A(0x18) -> V(0) N(0) Z(0) C(1) A(0x7E)\r
+\r
+H(0) C(0) A(0x19) -> V(0) N(0) Z(0) C(0) A(0x19)\r
+H(0) C(1) A(0x19) -> V(0) N(0) Z(0) C(1) A(0x79)\r
+H(1) C(0) A(0x19) -> V(0) N(0) Z(0) C(0) A(0x1F)\r
+H(1) C(1) A(0x19) -> V(0) N(0) Z(0) C(1) A(0x7F)\r
+\r
+H(0) C(0) A(0x1A) -> V(0) N(0) Z(0) C(0) A(0x20)\r
+H(0) C(1) A(0x1A) -> V(1) N(1) Z(0) C(1) A(0x80)\r
+H(1) C(0) A(0x1A) -> V(0) N(0) Z(0) C(0) A(0x20)\r
+H(1) C(1) A(0x1A) -> V(1) N(1) Z(0) C(1) A(0x80)\r
+\r
+H(0) C(0) A(0x1B) -> V(0) N(0) Z(0) C(0) A(0x21)\r
+H(0) C(1) A(0x1B) -> V(1) N(1) Z(0) C(1) A(0x81)\r
+H(1) C(0) A(0x1B) -> V(0) N(0) Z(0) C(0) A(0x21)\r
+H(1) C(1) A(0x1B) -> V(1) N(1) Z(0) C(1) A(0x81)\r
+\r
+H(0) C(0) A(0x1C) -> V(0) N(0) Z(0) C(0) A(0x22)\r
+H(0) C(1) A(0x1C) -> V(1) N(1) Z(0) C(1) A(0x82)\r
+H(1) C(0) A(0x1C) -> V(0) N(0) Z(0) C(0) A(0x22)\r
+H(1) C(1) A(0x1C) -> V(1) N(1) Z(0) C(1) A(0x82)\r
+\r
+H(0) C(0) A(0x1D) -> V(0) N(0) Z(0) C(0) A(0x23)\r
+H(0) C(1) A(0x1D) -> V(1) N(1) Z(0) C(1) A(0x83)\r
+H(1) C(0) A(0x1D) -> V(0) N(0) Z(0) C(0) A(0x23)\r
+H(1) C(1) A(0x1D) -> V(1) N(1) Z(0) C(1) A(0x83)\r
+\r
+H(0) C(0) A(0x1E) -> V(0) N(0) Z(0) C(0) A(0x24)\r
+H(0) C(1) A(0x1E) -> V(1) N(1) Z(0) C(1) A(0x84)\r
+H(1) C(0) A(0x1E) -> V(0) N(0) Z(0) C(0) A(0x24)\r
+H(1) C(1) A(0x1E) -> V(1) N(1) Z(0) C(1) A(0x84)\r
+\r
+H(0) C(0) A(0x1F) -> V(0) N(0) Z(0) C(0) A(0x25)\r
+H(0) C(1) A(0x1F) -> V(1) N(1) Z(0) C(1) A(0x85)\r
+H(1) C(0) A(0x1F) -> V(0) N(0) Z(0) C(0) A(0x25)\r
+H(1) C(1) A(0x1F) -> V(1) N(1) Z(0) C(1) A(0x85)\r
+\r
+H(0) C(0) A(0x20) -> V(0) N(0) Z(0) C(0) A(0x20)\r
+H(0) C(1) A(0x20) -> V(1) N(1) Z(0) C(1) A(0x80)\r
+H(1) C(0) A(0x20) -> V(0) N(0) Z(0) C(0) A(0x26)\r
+H(1) C(1) A(0x20) -> V(1) N(1) Z(0) C(1) A(0x86)\r
+\r
+H(0) C(0) A(0x21) -> V(0) N(0) Z(0) C(0) A(0x21)\r
+H(0) C(1) A(0x21) -> V(1) N(1) Z(0) C(1) A(0x81)\r
+H(1) C(0) A(0x21) -> V(0) N(0) Z(0) C(0) A(0x27)\r
+H(1) C(1) A(0x21) -> V(1) N(1) Z(0) C(1) A(0x87)\r
+\r
+H(0) C(0) A(0x22) -> V(0) N(0) Z(0) C(0) A(0x22)\r
+H(0) C(1) A(0x22) -> V(1) N(1) Z(0) C(1) A(0x82)\r
+H(1) C(0) A(0x22) -> V(0) N(0) Z(0) C(0) A(0x28)\r
+H(1) C(1) A(0x22) -> V(1) N(1) Z(0) C(1) A(0x88)\r
+\r
+H(0) C(0) A(0x23) -> V(0) N(0) Z(0) C(0) A(0x23)\r
+H(0) C(1) A(0x23) -> V(1) N(1) Z(0) C(1) A(0x83)\r
+H(1) C(0) A(0x23) -> V(0) N(0) Z(0) C(0) A(0x29)\r
+H(1) C(1) A(0x23) -> V(1) N(1) Z(0) C(1) A(0x89)\r
+\r
+H(0) C(0) A(0x24) -> V(0) N(0) Z(0) C(0) A(0x24)\r
+H(0) C(1) A(0x24) -> V(1) N(1) Z(0) C(1) A(0x84)\r
+H(1) C(0) A(0x24) -> V(0) N(0) Z(0) C(0) A(0x2A)\r
+H(1) C(1) A(0x24) -> V(1) N(1) Z(0) C(1) A(0x8A)\r
+\r
+H(0) C(0) A(0x25) -> V(0) N(0) Z(0) C(0) A(0x25)\r
+H(0) C(1) A(0x25) -> V(1) N(1) Z(0) C(1) A(0x85)\r
+H(1) C(0) A(0x25) -> V(0) N(0) Z(0) C(0) A(0x2B)\r
+H(1) C(1) A(0x25) -> V(1) N(1) Z(0) C(1) A(0x8B)\r
+\r
+H(0) C(0) A(0x26) -> V(0) N(0) Z(0) C(0) A(0x26)\r
+H(0) C(1) A(0x26) -> V(1) N(1) Z(0) C(1) A(0x86)\r
+H(1) C(0) A(0x26) -> V(0) N(0) Z(0) C(0) A(0x2C)\r
+H(1) C(1) A(0x26) -> V(1) N(1) Z(0) C(1) A(0x8C)\r
+\r
+H(0) C(0) A(0x27) -> V(0) N(0) Z(0) C(0) A(0x27)\r
+H(0) C(1) A(0x27) -> V(1) N(1) Z(0) C(1) A(0x87)\r
+H(1) C(0) A(0x27) -> V(0) N(0) Z(0) C(0) A(0x2D)\r
+H(1) C(1) A(0x27) -> V(1) N(1) Z(0) C(1) A(0x8D)\r
+\r
+H(0) C(0) A(0x28) -> V(0) N(0) Z(0) C(0) A(0x28)\r
+H(0) C(1) A(0x28) -> V(1) N(1) Z(0) C(1) A(0x88)\r
+H(1) C(0) A(0x28) -> V(0) N(0) Z(0) C(0) A(0x2E)\r
+H(1) C(1) A(0x28) -> V(1) N(1) Z(0) C(1) A(0x8E)\r
+\r
+H(0) C(0) A(0x29) -> V(0) N(0) Z(0) C(0) A(0x29)\r
+H(0) C(1) A(0x29) -> V(1) N(1) Z(0) C(1) A(0x89)\r
+H(1) C(0) A(0x29) -> V(0) N(0) Z(0) C(0) A(0x2F)\r
+H(1) C(1) A(0x29) -> V(1) N(1) Z(0) C(1) A(0x8F)\r
+\r
+H(0) C(0) A(0x2A) -> V(0) N(0) Z(0) C(0) A(0x30)\r
+H(0) C(1) A(0x2A) -> V(1) N(1) Z(0) C(1) A(0x90)\r
+H(1) C(0) A(0x2A) -> V(0) N(0) Z(0) C(0) A(0x30)\r
+H(1) C(1) A(0x2A) -> V(1) N(1) Z(0) C(1) A(0x90)\r
+\r
+H(0) C(0) A(0x2B) -> V(0) N(0) Z(0) C(0) A(0x31)\r
+H(0) C(1) A(0x2B) -> V(1) N(1) Z(0) C(1) A(0x91)\r
+H(1) C(0) A(0x2B) -> V(0) N(0) Z(0) C(0) A(0x31)\r
+H(1) C(1) A(0x2B) -> V(1) N(1) Z(0) C(1) A(0x91)\r
+\r
+H(0) C(0) A(0x2C) -> V(0) N(0) Z(0) C(0) A(0x32)\r
+H(0) C(1) A(0x2C) -> V(1) N(1) Z(0) C(1) A(0x92)\r
+H(1) C(0) A(0x2C) -> V(0) N(0) Z(0) C(0) A(0x32)\r
+H(1) C(1) A(0x2C) -> V(1) N(1) Z(0) C(1) A(0x92)\r
+\r
+H(0) C(0) A(0x2D) -> V(0) N(0) Z(0) C(0) A(0x33)\r
+H(0) C(1) A(0x2D) -> V(1) N(1) Z(0) C(1) A(0x93)\r
+H(1) C(0) A(0x2D) -> V(0) N(0) Z(0) C(0) A(0x33)\r
+H(1) C(1) A(0x2D) -> V(1) N(1) Z(0) C(1) A(0x93)\r
+\r
+H(0) C(0) A(0x2E) -> V(0) N(0) Z(0) C(0) A(0x34)\r
+H(0) C(1) A(0x2E) -> V(1) N(1) Z(0) C(1) A(0x94)\r
+H(1) C(0) A(0x2E) -> V(0) N(0) Z(0) C(0) A(0x34)\r
+H(1) C(1) A(0x2E) -> V(1) N(1) Z(0) C(1) A(0x94)\r
+\r
+H(0) C(0) A(0x2F) -> V(0) N(0) Z(0) C(0) A(0x35)\r
+H(0) C(1) A(0x2F) -> V(1) N(1) Z(0) C(1) A(0x95)\r
+H(1) C(0) A(0x2F) -> V(0) N(0) Z(0) C(0) A(0x35)\r
+H(1) C(1) A(0x2F) -> V(1) N(1) Z(0) C(1) A(0x95)\r
+\r
+H(0) C(0) A(0x30) -> V(0) N(0) Z(0) C(0) A(0x30)\r
+H(0) C(1) A(0x30) -> V(1) N(1) Z(0) C(1) A(0x90)\r
+H(1) C(0) A(0x30) -> V(0) N(0) Z(0) C(0) A(0x36)\r
+H(1) C(1) A(0x30) -> V(1) N(1) Z(0) C(1) A(0x96)\r
+\r
+H(0) C(0) A(0x31) -> V(0) N(0) Z(0) C(0) A(0x31)\r
+H(0) C(1) A(0x31) -> V(1) N(1) Z(0) C(1) A(0x91)\r
+H(1) C(0) A(0x31) -> V(0) N(0) Z(0) C(0) A(0x37)\r
+H(1) C(1) A(0x31) -> V(1) N(1) Z(0) C(1) A(0x97)\r
+\r
+H(0) C(0) A(0x32) -> V(0) N(0) Z(0) C(0) A(0x32)\r
+H(0) C(1) A(0x32) -> V(1) N(1) Z(0) C(1) A(0x92)\r
+H(1) C(0) A(0x32) -> V(0) N(0) Z(0) C(0) A(0x38)\r
+H(1) C(1) A(0x32) -> V(1) N(1) Z(0) C(1) A(0x98)\r
+\r
+H(0) C(0) A(0x33) -> V(0) N(0) Z(0) C(0) A(0x33)\r
+H(0) C(1) A(0x33) -> V(1) N(1) Z(0) C(1) A(0x93)\r
+H(1) C(0) A(0x33) -> V(0) N(0) Z(0) C(0) A(0x39)\r
+H(1) C(1) A(0x33) -> V(1) N(1) Z(0) C(1) A(0x99)\r
+\r
+H(0) C(0) A(0x34) -> V(0) N(0) Z(0) C(0) A(0x34)\r
+H(0) C(1) A(0x34) -> V(1) N(1) Z(0) C(1) A(0x94)\r
+H(1) C(0) A(0x34) -> V(0) N(0) Z(0) C(0) A(0x3A)\r
+H(1) C(1) A(0x34) -> V(1) N(1) Z(0) C(1) A(0x9A)\r
+\r
+H(0) C(0) A(0x35) -> V(0) N(0) Z(0) C(0) A(0x35)\r
+H(0) C(1) A(0x35) -> V(1) N(1) Z(0) C(1) A(0x95)\r
+H(1) C(0) A(0x35) -> V(0) N(0) Z(0) C(0) A(0x3B)\r
+H(1) C(1) A(0x35) -> V(1) N(1) Z(0) C(1) A(0x9B)\r
+\r
+H(0) C(0) A(0x36) -> V(0) N(0) Z(0) C(0) A(0x36)\r
+H(0) C(1) A(0x36) -> V(1) N(1) Z(0) C(1) A(0x96)\r
+H(1) C(0) A(0x36) -> V(0) N(0) Z(0) C(0) A(0x3C)\r
+H(1) C(1) A(0x36) -> V(1) N(1) Z(0) C(1) A(0x9C)\r
+\r
+H(0) C(0) A(0x37) -> V(0) N(0) Z(0) C(0) A(0x37)\r
+H(0) C(1) A(0x37) -> V(1) N(1) Z(0) C(1) A(0x97)\r
+H(1) C(0) A(0x37) -> V(0) N(0) Z(0) C(0) A(0x3D)\r
+H(1) C(1) A(0x37) -> V(1) N(1) Z(0) C(1) A(0x9D)\r
+\r
+H(0) C(0) A(0x38) -> V(0) N(0) Z(0) C(0) A(0x38)\r
+H(0) C(1) A(0x38) -> V(1) N(1) Z(0) C(1) A(0x98)\r
+H(1) C(0) A(0x38) -> V(0) N(0) Z(0) C(0) A(0x3E)\r
+H(1) C(1) A(0x38) -> V(1) N(1) Z(0) C(1) A(0x9E)\r
+\r
+H(0) C(0) A(0x39) -> V(0) N(0) Z(0) C(0) A(0x39)\r
+H(0) C(1) A(0x39) -> V(1) N(1) Z(0) C(1) A(0x99)\r
+H(1) C(0) A(0x39) -> V(0) N(0) Z(0) C(0) A(0x3F)\r
+H(1) C(1) A(0x39) -> V(1) N(1) Z(0) C(1) A(0x9F)\r
+\r
+H(0) C(0) A(0x3A) -> V(0) N(0) Z(0) C(0) A(0x40)\r
+H(0) C(1) A(0x3A) -> V(1) N(1) Z(0) C(1) A(0xA0)\r
+H(1) C(0) A(0x3A) -> V(0) N(0) Z(0) C(0) A(0x40)\r
+H(1) C(1) A(0x3A) -> V(1) N(1) Z(0) C(1) A(0xA0)\r
+\r
+H(0) C(0) A(0x3B) -> V(0) N(0) Z(0) C(0) A(0x41)\r
+H(0) C(1) A(0x3B) -> V(1) N(1) Z(0) C(1) A(0xA1)\r
+H(1) C(0) A(0x3B) -> V(0) N(0) Z(0) C(0) A(0x41)\r
+H(1) C(1) A(0x3B) -> V(1) N(1) Z(0) C(1) A(0xA1)\r
+\r
+H(0) C(0) A(0x3C) -> V(0) N(0) Z(0) C(0) A(0x42)\r
+H(0) C(1) A(0x3C) -> V(1) N(1) Z(0) C(1) A(0xA2)\r
+H(1) C(0) A(0x3C) -> V(0) N(0) Z(0) C(0) A(0x42)\r
+H(1) C(1) A(0x3C) -> V(1) N(1) Z(0) C(1) A(0xA2)\r
+\r
+H(0) C(0) A(0x3D) -> V(0) N(0) Z(0) C(0) A(0x43)\r
+H(0) C(1) A(0x3D) -> V(1) N(1) Z(0) C(1) A(0xA3)\r
+H(1) C(0) A(0x3D) -> V(0) N(0) Z(0) C(0) A(0x43)\r
+H(1) C(1) A(0x3D) -> V(1) N(1) Z(0) C(1) A(0xA3)\r
+\r
+H(0) C(0) A(0x3E) -> V(0) N(0) Z(0) C(0) A(0x44)\r
+H(0) C(1) A(0x3E) -> V(1) N(1) Z(0) C(1) A(0xA4)\r
+H(1) C(0) A(0x3E) -> V(0) N(0) Z(0) C(0) A(0x44)\r
+H(1) C(1) A(0x3E) -> V(1) N(1) Z(0) C(1) A(0xA4)\r
+\r
+H(0) C(0) A(0x3F) -> V(0) N(0) Z(0) C(0) A(0x45)\r
+H(0) C(1) A(0x3F) -> V(1) N(1) Z(0) C(1) A(0xA5)\r
+H(1) C(0) A(0x3F) -> V(0) N(0) Z(0) C(0) A(0x45)\r
+H(1) C(1) A(0x3F) -> V(1) N(1) Z(0) C(1) A(0xA5)\r
+\r
+H(0) C(0) A(0x40) -> V(0) N(0) Z(0) C(0) A(0x40)\r
+H(0) C(1) A(0x40) -> V(1) N(1) Z(0) C(1) A(0xA0)\r
+H(1) C(0) A(0x40) -> V(0) N(0) Z(0) C(0) A(0x46)\r
+H(1) C(1) A(0x40) -> V(1) N(1) Z(0) C(1) A(0xA6)\r
+\r
+H(0) C(0) A(0x41) -> V(0) N(0) Z(0) C(0) A(0x41)\r
+H(0) C(1) A(0x41) -> V(1) N(1) Z(0) C(1) A(0xA1)\r
+H(1) C(0) A(0x41) -> V(0) N(0) Z(0) C(0) A(0x47)\r
+H(1) C(1) A(0x41) -> V(1) N(1) Z(0) C(1) A(0xA7)\r
+\r
+H(0) C(0) A(0x42) -> V(0) N(0) Z(0) C(0) A(0x42)\r
+H(0) C(1) A(0x42) -> V(1) N(1) Z(0) C(1) A(0xA2)\r
+H(1) C(0) A(0x42) -> V(0) N(0) Z(0) C(0) A(0x48)\r
+H(1) C(1) A(0x42) -> V(1) N(1) Z(0) C(1) A(0xA8)\r
+\r
+H(0) C(0) A(0x43) -> V(0) N(0) Z(0) C(0) A(0x43)\r
+H(0) C(1) A(0x43) -> V(1) N(1) Z(0) C(1) A(0xA3)\r
+H(1) C(0) A(0x43) -> V(0) N(0) Z(0) C(0) A(0x49)\r
+H(1) C(1) A(0x43) -> V(1) N(1) Z(0) C(1) A(0xA9)\r
+\r
+H(0) C(0) A(0x44) -> V(0) N(0) Z(0) C(0) A(0x44)\r
+H(0) C(1) A(0x44) -> V(1) N(1) Z(0) C(1) A(0xA4)\r
+H(1) C(0) A(0x44) -> V(0) N(0) Z(0) C(0) A(0x4A)\r
+H(1) C(1) A(0x44) -> V(1) N(1) Z(0) C(1) A(0xAA)\r
+\r
+H(0) C(0) A(0x45) -> V(0) N(0) Z(0) C(0) A(0x45)\r
+H(0) C(1) A(0x45) -> V(1) N(1) Z(0) C(1) A(0xA5)\r
+H(1) C(0) A(0x45) -> V(0) N(0) Z(0) C(0) A(0x4B)\r
+H(1) C(1) A(0x45) -> V(1) N(1) Z(0) C(1) A(0xAB)\r
+\r
+H(0) C(0) A(0x46) -> V(0) N(0) Z(0) C(0) A(0x46)\r
+H(0) C(1) A(0x46) -> V(1) N(1) Z(0) C(1) A(0xA6)\r
+H(1) C(0) A(0x46) -> V(0) N(0) Z(0) C(0) A(0x4C)\r
+H(1) C(1) A(0x46) -> V(1) N(1) Z(0) C(1) A(0xAC)\r
+\r
+H(0) C(0) A(0x47) -> V(0) N(0) Z(0) C(0) A(0x47)\r
+H(0) C(1) A(0x47) -> V(1) N(1) Z(0) C(1) A(0xA7)\r
+H(1) C(0) A(0x47) -> V(0) N(0) Z(0) C(0) A(0x4D)\r
+H(1) C(1) A(0x47) -> V(1) N(1) Z(0) C(1) A(0xAD)\r
+\r
+H(0) C(0) A(0x48) -> V(0) N(0) Z(0) C(0) A(0x48)\r
+H(0) C(1) A(0x48) -> V(1) N(1) Z(0) C(1) A(0xA8)\r
+H(1) C(0) A(0x48) -> V(0) N(0) Z(0) C(0) A(0x4E)\r
+H(1) C(1) A(0x48) -> V(1) N(1) Z(0) C(1) A(0xAE)\r
+\r
+H(0) C(0) A(0x49) -> V(0) N(0) Z(0) C(0) A(0x49)\r
+H(0) C(1) A(0x49) -> V(1) N(1) Z(0) C(1) A(0xA9)\r
+H(1) C(0) A(0x49) -> V(0) N(0) Z(0) C(0) A(0x4F)\r
+H(1) C(1) A(0x49) -> V(1) N(1) Z(0) C(1) A(0xAF)\r
+\r
+H(0) C(0) A(0x4A) -> V(0) N(0) Z(0) C(0) A(0x50)\r
+H(0) C(1) A(0x4A) -> V(1) N(1) Z(0) C(1) A(0xB0)\r
+H(1) C(0) A(0x4A) -> V(0) N(0) Z(0) C(0) A(0x50)\r
+H(1) C(1) A(0x4A) -> V(1) N(1) Z(0) C(1) A(0xB0)\r
+\r
+H(0) C(0) A(0x4B) -> V(0) N(0) Z(0) C(0) A(0x51)\r
+H(0) C(1) A(0x4B) -> V(1) N(1) Z(0) C(1) A(0xB1)\r
+H(1) C(0) A(0x4B) -> V(0) N(0) Z(0) C(0) A(0x51)\r
+H(1) C(1) A(0x4B) -> V(1) N(1) Z(0) C(1) A(0xB1)\r
+\r
+H(0) C(0) A(0x4C) -> V(0) N(0) Z(0) C(0) A(0x52)\r
+H(0) C(1) A(0x4C) -> V(1) N(1) Z(0) C(1) A(0xB2)\r
+H(1) C(0) A(0x4C) -> V(0) N(0) Z(0) C(0) A(0x52)\r
+H(1) C(1) A(0x4C) -> V(1) N(1) Z(0) C(1) A(0xB2)\r
+\r
+H(0) C(0) A(0x4D) -> V(0) N(0) Z(0) C(0) A(0x53)\r
+H(0) C(1) A(0x4D) -> V(1) N(1) Z(0) C(1) A(0xB3)\r
+H(1) C(0) A(0x4D) -> V(0) N(0) Z(0) C(0) A(0x53)\r
+H(1) C(1) A(0x4D) -> V(1) N(1) Z(0) C(1) A(0xB3)\r
+\r
+H(0) C(0) A(0x4E) -> V(0) N(0) Z(0) C(0) A(0x54)\r
+H(0) C(1) A(0x4E) -> V(1) N(1) Z(0) C(1) A(0xB4)\r
+H(1) C(0) A(0x4E) -> V(0) N(0) Z(0) C(0) A(0x54)\r
+H(1) C(1) A(0x4E) -> V(1) N(1) Z(0) C(1) A(0xB4)\r
+\r
+H(0) C(0) A(0x4F) -> V(0) N(0) Z(0) C(0) A(0x55)\r
+H(0) C(1) A(0x4F) -> V(1) N(1) Z(0) C(1) A(0xB5)\r
+H(1) C(0) A(0x4F) -> V(0) N(0) Z(0) C(0) A(0x55)\r
+H(1) C(1) A(0x4F) -> V(1) N(1) Z(0) C(1) A(0xB5)\r
+\r
+H(0) C(0) A(0x50) -> V(0) N(0) Z(0) C(0) A(0x50)\r
+H(0) C(1) A(0x50) -> V(1) N(1) Z(0) C(1) A(0xB0)\r
+H(1) C(0) A(0x50) -> V(0) N(0) Z(0) C(0) A(0x56)\r
+H(1) C(1) A(0x50) -> V(1) N(1) Z(0) C(1) A(0xB6)\r
+\r
+H(0) C(0) A(0x51) -> V(0) N(0) Z(0) C(0) A(0x51)\r
+H(0) C(1) A(0x51) -> V(1) N(1) Z(0) C(1) A(0xB1)\r
+H(1) C(0) A(0x51) -> V(0) N(0) Z(0) C(0) A(0x57)\r
+H(1) C(1) A(0x51) -> V(1) N(1) Z(0) C(1) A(0xB7)\r
+\r
+H(0) C(0) A(0x52) -> V(0) N(0) Z(0) C(0) A(0x52)\r
+H(0) C(1) A(0x52) -> V(1) N(1) Z(0) C(1) A(0xB2)\r
+H(1) C(0) A(0x52) -> V(0) N(0) Z(0) C(0) A(0x58)\r
+H(1) C(1) A(0x52) -> V(1) N(1) Z(0) C(1) A(0xB8)\r
+\r
+H(0) C(0) A(0x53) -> V(0) N(0) Z(0) C(0) A(0x53)\r
+H(0) C(1) A(0x53) -> V(1) N(1) Z(0) C(1) A(0xB3)\r
+H(1) C(0) A(0x53) -> V(0) N(0) Z(0) C(0) A(0x59)\r
+H(1) C(1) A(0x53) -> V(1) N(1) Z(0) C(1) A(0xB9)\r
+\r
+H(0) C(0) A(0x54) -> V(0) N(0) Z(0) C(0) A(0x54)\r
+H(0) C(1) A(0x54) -> V(1) N(1) Z(0) C(1) A(0xB4)\r
+H(1) C(0) A(0x54) -> V(0) N(0) Z(0) C(0) A(0x5A)\r
+H(1) C(1) A(0x54) -> V(1) N(1) Z(0) C(1) A(0xBA)\r
+\r
+H(0) C(0) A(0x55) -> V(0) N(0) Z(0) C(0) A(0x55)\r
+H(0) C(1) A(0x55) -> V(1) N(1) Z(0) C(1) A(0xB5)\r
+H(1) C(0) A(0x55) -> V(0) N(0) Z(0) C(0) A(0x5B)\r
+H(1) C(1) A(0x55) -> V(1) N(1) Z(0) C(1) A(0xBB)\r
+\r
+H(0) C(0) A(0x56) -> V(0) N(0) Z(0) C(0) A(0x56)\r
+H(0) C(1) A(0x56) -> V(1) N(1) Z(0) C(1) A(0xB6)\r
+H(1) C(0) A(0x56) -> V(0) N(0) Z(0) C(0) A(0x5C)\r
+H(1) C(1) A(0x56) -> V(1) N(1) Z(0) C(1) A(0xBC)\r
+\r
+H(0) C(0) A(0x57) -> V(0) N(0) Z(0) C(0) A(0x57)\r
+H(0) C(1) A(0x57) -> V(1) N(1) Z(0) C(1) A(0xB7)\r
+H(1) C(0) A(0x57) -> V(0) N(0) Z(0) C(0) A(0x5D)\r
+H(1) C(1) A(0x57) -> V(1) N(1) Z(0) C(1) A(0xBD)\r
+\r
+H(0) C(0) A(0x58) -> V(0) N(0) Z(0) C(0) A(0x58)\r
+H(0) C(1) A(0x58) -> V(1) N(1) Z(0) C(1) A(0xB8)\r
+H(1) C(0) A(0x58) -> V(0) N(0) Z(0) C(0) A(0x5E)\r
+H(1) C(1) A(0x58) -> V(1) N(1) Z(0) C(1) A(0xBE)\r
+\r
+H(0) C(0) A(0x59) -> V(0) N(0) Z(0) C(0) A(0x59)\r
+H(0) C(1) A(0x59) -> V(1) N(1) Z(0) C(1) A(0xB9)\r
+H(1) C(0) A(0x59) -> V(0) N(0) Z(0) C(0) A(0x5F)\r
+H(1) C(1) A(0x59) -> V(1) N(1) Z(0) C(1) A(0xBF)\r
+\r
+H(0) C(0) A(0x5A) -> V(0) N(0) Z(0) C(0) A(0x60)\r
+H(0) C(1) A(0x5A) -> V(1) N(1) Z(0) C(1) A(0xC0)\r
+H(1) C(0) A(0x5A) -> V(0) N(0) Z(0) C(0) A(0x60)\r
+H(1) C(1) A(0x5A) -> V(1) N(1) Z(0) C(1) A(0xC0)\r
+\r
+H(0) C(0) A(0x5B) -> V(0) N(0) Z(0) C(0) A(0x61)\r
+H(0) C(1) A(0x5B) -> V(1) N(1) Z(0) C(1) A(0xC1)\r
+H(1) C(0) A(0x5B) -> V(0) N(0) Z(0) C(0) A(0x61)\r
+H(1) C(1) A(0x5B) -> V(1) N(1) Z(0) C(1) A(0xC1)\r
+\r
+H(0) C(0) A(0x5C) -> V(0) N(0) Z(0) C(0) A(0x62)\r
+H(0) C(1) A(0x5C) -> V(1) N(1) Z(0) C(1) A(0xC2)\r
+H(1) C(0) A(0x5C) -> V(0) N(0) Z(0) C(0) A(0x62)\r
+H(1) C(1) A(0x5C) -> V(1) N(1) Z(0) C(1) A(0xC2)\r
+\r
+H(0) C(0) A(0x5D) -> V(0) N(0) Z(0) C(0) A(0x63)\r
+H(0) C(1) A(0x5D) -> V(1) N(1) Z(0) C(1) A(0xC3)\r
+H(1) C(0) A(0x5D) -> V(0) N(0) Z(0) C(0) A(0x63)\r
+H(1) C(1) A(0x5D) -> V(1) N(1) Z(0) C(1) A(0xC3)\r
+\r
+H(0) C(0) A(0x5E) -> V(0) N(0) Z(0) C(0) A(0x64)\r
+H(0) C(1) A(0x5E) -> V(1) N(1) Z(0) C(1) A(0xC4)\r
+H(1) C(0) A(0x5E) -> V(0) N(0) Z(0) C(0) A(0x64)\r
+H(1) C(1) A(0x5E) -> V(1) N(1) Z(0) C(1) A(0xC4)\r
+\r
+H(0) C(0) A(0x5F) -> V(0) N(0) Z(0) C(0) A(0x65)\r
+H(0) C(1) A(0x5F) -> V(1) N(1) Z(0) C(1) A(0xC5)\r
+H(1) C(0) A(0x5F) -> V(0) N(0) Z(0) C(0) A(0x65)\r
+H(1) C(1) A(0x5F) -> V(1) N(1) Z(0) C(1) A(0xC5)\r
+\r
+H(0) C(0) A(0x60) -> V(0) N(0) Z(0) C(0) A(0x60)\r
+H(0) C(1) A(0x60) -> V(1) N(1) Z(0) C(1) A(0xC0)\r
+H(1) C(0) A(0x60) -> V(0) N(0) Z(0) C(0) A(0x66)\r
+H(1) C(1) A(0x60) -> V(1) N(1) Z(0) C(1) A(0xC6)\r
+\r
+H(0) C(0) A(0x61) -> V(0) N(0) Z(0) C(0) A(0x61)\r
+H(0) C(1) A(0x61) -> V(1) N(1) Z(0) C(1) A(0xC1)\r
+H(1) C(0) A(0x61) -> V(0) N(0) Z(0) C(0) A(0x67)\r
+H(1) C(1) A(0x61) -> V(1) N(1) Z(0) C(1) A(0xC7)\r
+\r
+H(0) C(0) A(0x62) -> V(0) N(0) Z(0) C(0) A(0x62)\r
+H(0) C(1) A(0x62) -> V(1) N(1) Z(0) C(1) A(0xC2)\r
+H(1) C(0) A(0x62) -> V(0) N(0) Z(0) C(0) A(0x68)\r
+H(1) C(1) A(0x62) -> V(1) N(1) Z(0) C(1) A(0xC8)\r
+\r
+H(0) C(0) A(0x63) -> V(0) N(0) Z(0) C(0) A(0x63)\r
+H(0) C(1) A(0x63) -> V(1) N(1) Z(0) C(1) A(0xC3)\r
+H(1) C(0) A(0x63) -> V(0) N(0) Z(0) C(0) A(0x69)\r
+H(1) C(1) A(0x63) -> V(1) N(1) Z(0) C(1) A(0xC9)\r
+\r
+H(0) C(0) A(0x64) -> V(0) N(0) Z(0) C(0) A(0x64)\r
+H(0) C(1) A(0x64) -> V(1) N(1) Z(0) C(1) A(0xC4)\r
+H(1) C(0) A(0x64) -> V(0) N(0) Z(0) C(0) A(0x6A)\r
+H(1) C(1) A(0x64) -> V(1) N(1) Z(0) C(1) A(0xCA)\r
+\r
+H(0) C(0) A(0x65) -> V(0) N(0) Z(0) C(0) A(0x65)\r
+H(0) C(1) A(0x65) -> V(1) N(1) Z(0) C(1) A(0xC5)\r
+H(1) C(0) A(0x65) -> V(0) N(0) Z(0) C(0) A(0x6B)\r
+H(1) C(1) A(0x65) -> V(1) N(1) Z(0) C(1) A(0xCB)\r
+\r
+H(0) C(0) A(0x66) -> V(0) N(0) Z(0) C(0) A(0x66)\r
+H(0) C(1) A(0x66) -> V(1) N(1) Z(0) C(1) A(0xC6)\r
+H(1) C(0) A(0x66) -> V(0) N(0) Z(0) C(0) A(0x6C)\r
+H(1) C(1) A(0x66) -> V(1) N(1) Z(0) C(1) A(0xCC)\r
+\r
+H(0) C(0) A(0x67) -> V(0) N(0) Z(0) C(0) A(0x67)\r
+H(0) C(1) A(0x67) -> V(1) N(1) Z(0) C(1) A(0xC7)\r
+H(1) C(0) A(0x67) -> V(0) N(0) Z(0) C(0) A(0x6D)\r
+H(1) C(1) A(0x67) -> V(1) N(1) Z(0) C(1) A(0xCD)\r
+\r
+H(0) C(0) A(0x68) -> V(0) N(0) Z(0) C(0) A(0x68)\r
+H(0) C(1) A(0x68) -> V(1) N(1) Z(0) C(1) A(0xC8)\r
+H(1) C(0) A(0x68) -> V(0) N(0) Z(0) C(0) A(0x6E)\r
+H(1) C(1) A(0x68) -> V(1) N(1) Z(0) C(1) A(0xCE)\r
+\r
+H(0) C(0) A(0x69) -> V(0) N(0) Z(0) C(0) A(0x69)\r
+H(0) C(1) A(0x69) -> V(1) N(1) Z(0) C(1) A(0xC9)\r
+H(1) C(0) A(0x69) -> V(0) N(0) Z(0) C(0) A(0x6F)\r
+H(1) C(1) A(0x69) -> V(1) N(1) Z(0) C(1) A(0xCF)\r
+\r
+H(0) C(0) A(0x6A) -> V(0) N(0) Z(0) C(0) A(0x70)\r
+H(0) C(1) A(0x6A) -> V(1) N(1) Z(0) C(1) A(0xD0)\r
+H(1) C(0) A(0x6A) -> V(0) N(0) Z(0) C(0) A(0x70)\r
+H(1) C(1) A(0x6A) -> V(1) N(1) Z(0) C(1) A(0xD0)\r
+\r
+H(0) C(0) A(0x6B) -> V(0) N(0) Z(0) C(0) A(0x71)\r
+H(0) C(1) A(0x6B) -> V(1) N(1) Z(0) C(1) A(0xD1)\r
+H(1) C(0) A(0x6B) -> V(0) N(0) Z(0) C(0) A(0x71)\r
+H(1) C(1) A(0x6B) -> V(1) N(1) Z(0) C(1) A(0xD1)\r
+\r
+H(0) C(0) A(0x6C) -> V(0) N(0) Z(0) C(0) A(0x72)\r
+H(0) C(1) A(0x6C) -> V(1) N(1) Z(0) C(1) A(0xD2)\r
+H(1) C(0) A(0x6C) -> V(0) N(0) Z(0) C(0) A(0x72)\r
+H(1) C(1) A(0x6C) -> V(1) N(1) Z(0) C(1) A(0xD2)\r
+\r
+H(0) C(0) A(0x6D) -> V(0) N(0) Z(0) C(0) A(0x73)\r
+H(0) C(1) A(0x6D) -> V(1) N(1) Z(0) C(1) A(0xD3)\r
+H(1) C(0) A(0x6D) -> V(0) N(0) Z(0) C(0) A(0x73)\r
+H(1) C(1) A(0x6D) -> V(1) N(1) Z(0) C(1) A(0xD3)\r
+\r
+H(0) C(0) A(0x6E) -> V(0) N(0) Z(0) C(0) A(0x74)\r
+H(0) C(1) A(0x6E) -> V(1) N(1) Z(0) C(1) A(0xD4)\r
+H(1) C(0) A(0x6E) -> V(0) N(0) Z(0) C(0) A(0x74)\r
+H(1) C(1) A(0x6E) -> V(1) N(1) Z(0) C(1) A(0xD4)\r
+\r
+H(0) C(0) A(0x6F) -> V(0) N(0) Z(0) C(0) A(0x75)\r
+H(0) C(1) A(0x6F) -> V(1) N(1) Z(0) C(1) A(0xD5)\r
+H(1) C(0) A(0x6F) -> V(0) N(0) Z(0) C(0) A(0x75)\r
+H(1) C(1) A(0x6F) -> V(1) N(1) Z(0) C(1) A(0xD5)\r
+\r
+H(0) C(0) A(0x70) -> V(0) N(0) Z(0) C(0) A(0x70)\r
+H(0) C(1) A(0x70) -> V(1) N(1) Z(0) C(1) A(0xD0)\r
+H(1) C(0) A(0x70) -> V(0) N(0) Z(0) C(0) A(0x76)\r
+H(1) C(1) A(0x70) -> V(1) N(1) Z(0) C(1) A(0xD6)\r
+\r
+H(0) C(0) A(0x71) -> V(0) N(0) Z(0) C(0) A(0x71)\r
+H(0) C(1) A(0x71) -> V(1) N(1) Z(0) C(1) A(0xD1)\r
+H(1) C(0) A(0x71) -> V(0) N(0) Z(0) C(0) A(0x77)\r
+H(1) C(1) A(0x71) -> V(1) N(1) Z(0) C(1) A(0xD7)\r
+\r
+H(0) C(0) A(0x72) -> V(0) N(0) Z(0) C(0) A(0x72)\r
+H(0) C(1) A(0x72) -> V(1) N(1) Z(0) C(1) A(0xD2)\r
+H(1) C(0) A(0x72) -> V(0) N(0) Z(0) C(0) A(0x78)\r
+H(1) C(1) A(0x72) -> V(1) N(1) Z(0) C(1) A(0xD8)\r
+\r
+H(0) C(0) A(0x73) -> V(0) N(0) Z(0) C(0) A(0x73)\r
+H(0) C(1) A(0x73) -> V(1) N(1) Z(0) C(1) A(0xD3)\r
+H(1) C(0) A(0x73) -> V(0) N(0) Z(0) C(0) A(0x79)\r
+H(1) C(1) A(0x73) -> V(1) N(1) Z(0) C(1) A(0xD9)\r
+\r
+H(0) C(0) A(0x74) -> V(0) N(0) Z(0) C(0) A(0x74)\r
+H(0) C(1) A(0x74) -> V(1) N(1) Z(0) C(1) A(0xD4)\r
+H(1) C(0) A(0x74) -> V(0) N(0) Z(0) C(0) A(0x7A)\r
+H(1) C(1) A(0x74) -> V(1) N(1) Z(0) C(1) A(0xDA)\r
+\r
+H(0) C(0) A(0x75) -> V(0) N(0) Z(0) C(0) A(0x75)\r
+H(0) C(1) A(0x75) -> V(1) N(1) Z(0) C(1) A(0xD5)\r
+H(1) C(0) A(0x75) -> V(0) N(0) Z(0) C(0) A(0x7B)\r
+H(1) C(1) A(0x75) -> V(1) N(1) Z(0) C(1) A(0xDB)\r
+\r
+H(0) C(0) A(0x76) -> V(0) N(0) Z(0) C(0) A(0x76)\r
+H(0) C(1) A(0x76) -> V(1) N(1) Z(0) C(1) A(0xD6)\r
+H(1) C(0) A(0x76) -> V(0) N(0) Z(0) C(0) A(0x7C)\r
+H(1) C(1) A(0x76) -> V(1) N(1) Z(0) C(1) A(0xDC)\r
+\r
+H(0) C(0) A(0x77) -> V(0) N(0) Z(0) C(0) A(0x77)\r
+H(0) C(1) A(0x77) -> V(1) N(1) Z(0) C(1) A(0xD7)\r
+H(1) C(0) A(0x77) -> V(0) N(0) Z(0) C(0) A(0x7D)\r
+H(1) C(1) A(0x77) -> V(1) N(1) Z(0) C(1) A(0xDD)\r
+\r
+H(0) C(0) A(0x78) -> V(0) N(0) Z(0) C(0) A(0x78)\r
+H(0) C(1) A(0x78) -> V(1) N(1) Z(0) C(1) A(0xD8)\r
+H(1) C(0) A(0x78) -> V(0) N(0) Z(0) C(0) A(0x7E)\r
+H(1) C(1) A(0x78) -> V(1) N(1) Z(0) C(1) A(0xDE)\r
+\r
+H(0) C(0) A(0x79) -> V(0) N(0) Z(0) C(0) A(0x79)\r
+H(0) C(1) A(0x79) -> V(1) N(1) Z(0) C(1) A(0xD9)\r
+H(1) C(0) A(0x79) -> V(0) N(0) Z(0) C(0) A(0x7F)\r
+H(1) C(1) A(0x79) -> V(1) N(1) Z(0) C(1) A(0xDF)\r
+\r
+H(0) C(0) A(0x7A) -> V(1) N(1) Z(0) C(0) A(0x80)\r
+H(0) C(1) A(0x7A) -> V(1) N(1) Z(0) C(1) A(0xE0)\r
+H(1) C(0) A(0x7A) -> V(1) N(1) Z(0) C(0) A(0x80)\r
+H(1) C(1) A(0x7A) -> V(1) N(1) Z(0) C(1) A(0xE0)\r
+\r
+H(0) C(0) A(0x7B) -> V(1) N(1) Z(0) C(0) A(0x81)\r
+H(0) C(1) A(0x7B) -> V(1) N(1) Z(0) C(1) A(0xE1)\r
+H(1) C(0) A(0x7B) -> V(1) N(1) Z(0) C(0) A(0x81)\r
+H(1) C(1) A(0x7B) -> V(1) N(1) Z(0) C(1) A(0xE1)\r
+\r
+H(0) C(0) A(0x7C) -> V(1) N(1) Z(0) C(0) A(0x82)\r
+H(0) C(1) A(0x7C) -> V(1) N(1) Z(0) C(1) A(0xE2)\r
+H(1) C(0) A(0x7C) -> V(1) N(1) Z(0) C(0) A(0x82)\r
+H(1) C(1) A(0x7C) -> V(1) N(1) Z(0) C(1) A(0xE2)\r
+\r
+H(0) C(0) A(0x7D) -> V(1) N(1) Z(0) C(0) A(0x83)\r
+H(0) C(1) A(0x7D) -> V(1) N(1) Z(0) C(1) A(0xE3)\r
+H(1) C(0) A(0x7D) -> V(1) N(1) Z(0) C(0) A(0x83)\r
+H(1) C(1) A(0x7D) -> V(1) N(1) Z(0) C(1) A(0xE3)\r
+\r
+H(0) C(0) A(0x7E) -> V(1) N(1) Z(0) C(0) A(0x84)\r
+H(0) C(1) A(0x7E) -> V(1) N(1) Z(0) C(1) A(0xE4)\r
+H(1) C(0) A(0x7E) -> V(1) N(1) Z(0) C(0) A(0x84)\r
+H(1) C(1) A(0x7E) -> V(1) N(1) Z(0) C(1) A(0xE4)\r
+\r
+H(0) C(0) A(0x7F) -> V(1) N(1) Z(0) C(0) A(0x85)\r
+H(0) C(1) A(0x7F) -> V(1) N(1) Z(0) C(1) A(0xE5)\r
+H(1) C(0) A(0x7F) -> V(1) N(1) Z(0) C(0) A(0x85)\r
+H(1) C(1) A(0x7F) -> V(1) N(1) Z(0) C(1) A(0xE5)\r
+\r
+H(0) C(0) A(0x80) -> V(0) N(1) Z(0) C(0) A(0x80)\r
+H(0) C(1) A(0x80) -> V(0) N(1) Z(0) C(1) A(0xE0)\r
+H(1) C(0) A(0x80) -> V(0) N(1) Z(0) C(0) A(0x86)\r
+H(1) C(1) A(0x80) -> V(0) N(1) Z(0) C(1) A(0xE6)\r
+\r
+H(0) C(0) A(0x81) -> V(0) N(1) Z(0) C(0) A(0x81)\r
+H(0) C(1) A(0x81) -> V(0) N(1) Z(0) C(1) A(0xE1)\r
+H(1) C(0) A(0x81) -> V(0) N(1) Z(0) C(0) A(0x87)\r
+H(1) C(1) A(0x81) -> V(0) N(1) Z(0) C(1) A(0xE7)\r
+\r
+H(0) C(0) A(0x82) -> V(0) N(1) Z(0) C(0) A(0x82)\r
+H(0) C(1) A(0x82) -> V(0) N(1) Z(0) C(1) A(0xE2)\r
+H(1) C(0) A(0x82) -> V(0) N(1) Z(0) C(0) A(0x88)\r
+H(1) C(1) A(0x82) -> V(0) N(1) Z(0) C(1) A(0xE8)\r
+\r
+H(0) C(0) A(0x83) -> V(0) N(1) Z(0) C(0) A(0x83)\r
+H(0) C(1) A(0x83) -> V(0) N(1) Z(0) C(1) A(0xE3)\r
+H(1) C(0) A(0x83) -> V(0) N(1) Z(0) C(0) A(0x89)\r
+H(1) C(1) A(0x83) -> V(0) N(1) Z(0) C(1) A(0xE9)\r
+\r
+H(0) C(0) A(0x84) -> V(0) N(1) Z(0) C(0) A(0x84)\r
+H(0) C(1) A(0x84) -> V(0) N(1) Z(0) C(1) A(0xE4)\r
+H(1) C(0) A(0x84) -> V(0) N(1) Z(0) C(0) A(0x8A)\r
+H(1) C(1) A(0x84) -> V(0) N(1) Z(0) C(1) A(0xEA)\r
+\r
+H(0) C(0) A(0x85) -> V(0) N(1) Z(0) C(0) A(0x85)\r
+H(0) C(1) A(0x85) -> V(0) N(1) Z(0) C(1) A(0xE5)\r
+H(1) C(0) A(0x85) -> V(0) N(1) Z(0) C(0) A(0x8B)\r
+H(1) C(1) A(0x85) -> V(0) N(1) Z(0) C(1) A(0xEB)\r
+\r
+H(0) C(0) A(0x86) -> V(0) N(1) Z(0) C(0) A(0x86)\r
+H(0) C(1) A(0x86) -> V(0) N(1) Z(0) C(1) A(0xE6)\r
+H(1) C(0) A(0x86) -> V(0) N(1) Z(0) C(0) A(0x8C)\r
+H(1) C(1) A(0x86) -> V(0) N(1) Z(0) C(1) A(0xEC)\r
+\r
+H(0) C(0) A(0x87) -> V(0) N(1) Z(0) C(0) A(0x87)\r
+H(0) C(1) A(0x87) -> V(0) N(1) Z(0) C(1) A(0xE7)\r
+H(1) C(0) A(0x87) -> V(0) N(1) Z(0) C(0) A(0x8D)\r
+H(1) C(1) A(0x87) -> V(0) N(1) Z(0) C(1) A(0xED)\r
+\r
+H(0) C(0) A(0x88) -> V(0) N(1) Z(0) C(0) A(0x88)\r
+H(0) C(1) A(0x88) -> V(0) N(1) Z(0) C(1) A(0xE8)\r
+H(1) C(0) A(0x88) -> V(0) N(1) Z(0) C(0) A(0x8E)\r
+H(1) C(1) A(0x88) -> V(0) N(1) Z(0) C(1) A(0xEE)\r
+\r
+H(0) C(0) A(0x89) -> V(0) N(1) Z(0) C(0) A(0x89)\r
+H(0) C(1) A(0x89) -> V(0) N(1) Z(0) C(1) A(0xE9)\r
+H(1) C(0) A(0x89) -> V(0) N(1) Z(0) C(0) A(0x8F)\r
+H(1) C(1) A(0x89) -> V(0) N(1) Z(0) C(1) A(0xEF)\r
+\r
+H(0) C(0) A(0x8A) -> V(0) N(1) Z(0) C(0) A(0x90)\r
+H(0) C(1) A(0x8A) -> V(0) N(1) Z(0) C(1) A(0xF0)\r
+H(1) C(0) A(0x8A) -> V(0) N(1) Z(0) C(0) A(0x90)\r
+H(1) C(1) A(0x8A) -> V(0) N(1) Z(0) C(1) A(0xF0)\r
+\r
+H(0) C(0) A(0x8B) -> V(0) N(1) Z(0) C(0) A(0x91)\r
+H(0) C(1) A(0x8B) -> V(0) N(1) Z(0) C(1) A(0xF1)\r
+H(1) C(0) A(0x8B) -> V(0) N(1) Z(0) C(0) A(0x91)\r
+H(1) C(1) A(0x8B) -> V(0) N(1) Z(0) C(1) A(0xF1)\r
+\r
+H(0) C(0) A(0x8C) -> V(0) N(1) Z(0) C(0) A(0x92)\r
+H(0) C(1) A(0x8C) -> V(0) N(1) Z(0) C(1) A(0xF2)\r
+H(1) C(0) A(0x8C) -> V(0) N(1) Z(0) C(0) A(0x92)\r
+H(1) C(1) A(0x8C) -> V(0) N(1) Z(0) C(1) A(0xF2)\r
+\r
+H(0) C(0) A(0x8D) -> V(0) N(1) Z(0) C(0) A(0x93)\r
+H(0) C(1) A(0x8D) -> V(0) N(1) Z(0) C(1) A(0xF3)\r
+H(1) C(0) A(0x8D) -> V(0) N(1) Z(0) C(0) A(0x93)\r
+H(1) C(1) A(0x8D) -> V(0) N(1) Z(0) C(1) A(0xF3)\r
+\r
+H(0) C(0) A(0x8E) -> V(0) N(1) Z(0) C(0) A(0x94)\r
+H(0) C(1) A(0x8E) -> V(0) N(1) Z(0) C(1) A(0xF4)\r
+H(1) C(0) A(0x8E) -> V(0) N(1) Z(0) C(0) A(0x94)\r
+H(1) C(1) A(0x8E) -> V(0) N(1) Z(0) C(1) A(0xF4)\r
+\r
+H(0) C(0) A(0x8F) -> V(0) N(1) Z(0) C(0) A(0x95)\r
+H(0) C(1) A(0x8F) -> V(0) N(1) Z(0) C(1) A(0xF5)\r
+H(1) C(0) A(0x8F) -> V(0) N(1) Z(0) C(0) A(0x95)\r
+H(1) C(1) A(0x8F) -> V(0) N(1) Z(0) C(1) A(0xF5)\r
+\r
+H(0) C(0) A(0x90) -> V(0) N(1) Z(0) C(0) A(0x90)\r
+H(0) C(1) A(0x90) -> V(0) N(1) Z(0) C(1) A(0xF0)\r
+H(1) C(0) A(0x90) -> V(0) N(1) Z(0) C(0) A(0x96)\r
+H(1) C(1) A(0x90) -> V(0) N(1) Z(0) C(1) A(0xF6)\r
+\r
+H(0) C(0) A(0x91) -> V(0) N(1) Z(0) C(0) A(0x91)\r
+H(0) C(1) A(0x91) -> V(0) N(1) Z(0) C(1) A(0xF1)\r
+H(1) C(0) A(0x91) -> V(0) N(1) Z(0) C(0) A(0x97)\r
+H(1) C(1) A(0x91) -> V(0) N(1) Z(0) C(1) A(0xF7)\r
+\r
+H(0) C(0) A(0x92) -> V(0) N(1) Z(0) C(0) A(0x92)\r
+H(0) C(1) A(0x92) -> V(0) N(1) Z(0) C(1) A(0xF2)\r
+H(1) C(0) A(0x92) -> V(0) N(1) Z(0) C(0) A(0x98)\r
+H(1) C(1) A(0x92) -> V(0) N(1) Z(0) C(1) A(0xF8)\r
+\r
+H(0) C(0) A(0x93) -> V(0) N(1) Z(0) C(0) A(0x93)\r
+H(0) C(1) A(0x93) -> V(0) N(1) Z(0) C(1) A(0xF3)\r
+H(1) C(0) A(0x93) -> V(0) N(1) Z(0) C(0) A(0x99)\r
+H(1) C(1) A(0x93) -> V(0) N(1) Z(0) C(1) A(0xF9)\r
+\r
+H(0) C(0) A(0x94) -> V(0) N(1) Z(0) C(0) A(0x94)\r
+H(0) C(1) A(0x94) -> V(0) N(1) Z(0) C(1) A(0xF4)\r
+H(1) C(0) A(0x94) -> V(0) N(1) Z(0) C(0) A(0x9A)\r
+H(1) C(1) A(0x94) -> V(0) N(1) Z(0) C(1) A(0xFA)\r
+\r
+H(0) C(0) A(0x95) -> V(0) N(1) Z(0) C(0) A(0x95)\r
+H(0) C(1) A(0x95) -> V(0) N(1) Z(0) C(1) A(0xF5)\r
+H(1) C(0) A(0x95) -> V(0) N(1) Z(0) C(0) A(0x9B)\r
+H(1) C(1) A(0x95) -> V(0) N(1) Z(0) C(1) A(0xFB)\r
+\r
+H(0) C(0) A(0x96) -> V(0) N(1) Z(0) C(0) A(0x96)\r
+H(0) C(1) A(0x96) -> V(0) N(1) Z(0) C(1) A(0xF6)\r
+H(1) C(0) A(0x96) -> V(0) N(1) Z(0) C(0) A(0x9C)\r
+H(1) C(1) A(0x96) -> V(0) N(1) Z(0) C(1) A(0xFC)\r
+\r
+H(0) C(0) A(0x97) -> V(0) N(1) Z(0) C(0) A(0x97)\r
+H(0) C(1) A(0x97) -> V(0) N(1) Z(0) C(1) A(0xF7)\r
+H(1) C(0) A(0x97) -> V(0) N(1) Z(0) C(0) A(0x9D)\r
+H(1) C(1) A(0x97) -> V(0) N(1) Z(0) C(1) A(0xFD)\r
+\r
+H(0) C(0) A(0x98) -> V(0) N(1) Z(0) C(0) A(0x98)\r
+H(0) C(1) A(0x98) -> V(0) N(1) Z(0) C(1) A(0xF8)\r
+H(1) C(0) A(0x98) -> V(0) N(1) Z(0) C(0) A(0x9E)\r
+H(1) C(1) A(0x98) -> V(0) N(1) Z(0) C(1) A(0xFE)\r
+\r
+H(0) C(0) A(0x99) -> V(0) N(1) Z(0) C(0) A(0x99)\r
+H(0) C(1) A(0x99) -> V(0) N(1) Z(0) C(1) A(0xF9)\r
+H(1) C(0) A(0x99) -> V(0) N(1) Z(0) C(0) A(0x9F)\r
+H(1) C(1) A(0x99) -> V(0) N(1) Z(0) C(1) A(0xFF)\r
+\r
+H(0) C(0) A(0x9A) -> V(0) N(0) Z(1) C(1) A(0x00)\r
+H(0) C(1) A(0x9A) -> V(0) N(0) Z(1) C(1) A(0x00)\r
+H(1) C(0) A(0x9A) -> V(0) N(0) Z(1) C(1) A(0x00)\r
+H(1) C(1) A(0x9A) -> V(0) N(0) Z(1) C(1) A(0x00)\r
+\r
+H(0) C(0) A(0x9B) -> V(0) N(0) Z(0) C(1) A(0x01)\r
+H(0) C(1) A(0x9B) -> V(0) N(0) Z(0) C(1) A(0x01)\r
+H(1) C(0) A(0x9B) -> V(0) N(0) Z(0) C(1) A(0x01)\r
+H(1) C(1) A(0x9B) -> V(0) N(0) Z(0) C(1) A(0x01)\r
+\r
+H(0) C(0) A(0x9C) -> V(0) N(0) Z(0) C(1) A(0x02)\r
+H(0) C(1) A(0x9C) -> V(0) N(0) Z(0) C(1) A(0x02)\r
+H(1) C(0) A(0x9C) -> V(0) N(0) Z(0) C(1) A(0x02)\r
+H(1) C(1) A(0x9C) -> V(0) N(0) Z(0) C(1) A(0x02)\r
+\r
+H(0) C(0) A(0x9D) -> V(0) N(0) Z(0) C(1) A(0x03)\r
+H(0) C(1) A(0x9D) -> V(0) N(0) Z(0) C(1) A(0x03)\r
+H(1) C(0) A(0x9D) -> V(0) N(0) Z(0) C(1) A(0x03)\r
+H(1) C(1) A(0x9D) -> V(0) N(0) Z(0) C(1) A(0x03)\r
+\r
+H(0) C(0) A(0x9E) -> V(0) N(0) Z(0) C(1) A(0x04)\r
+H(0) C(1) A(0x9E) -> V(0) N(0) Z(0) C(1) A(0x04)\r
+H(1) C(0) A(0x9E) -> V(0) N(0) Z(0) C(1) A(0x04)\r
+H(1) C(1) A(0x9E) -> V(0) N(0) Z(0) C(1) A(0x04)\r
+\r
+H(0) C(0) A(0x9F) -> V(0) N(0) Z(0) C(1) A(0x05)\r
+H(0) C(1) A(0x9F) -> V(0) N(0) Z(0) C(1) A(0x05)\r
+H(1) C(0) A(0x9F) -> V(0) N(0) Z(0) C(1) A(0x05)\r
+H(1) C(1) A(0x9F) -> V(0) N(0) Z(0) C(1) A(0x05)\r
+\r
+H(0) C(0) A(0xA0) -> V(0) N(0) Z(1) C(1) A(0x00)\r
+H(0) C(1) A(0xA0) -> V(0) N(0) Z(1) C(1) A(0x00)\r
+H(1) C(0) A(0xA0) -> V(0) N(0) Z(0) C(1) A(0x06)\r
+H(1) C(1) A(0xA0) -> V(0) N(0) Z(0) C(1) A(0x06)\r
+\r
+H(0) C(0) A(0xA1) -> V(0) N(0) Z(0) C(1) A(0x01)\r
+H(0) C(1) A(0xA1) -> V(0) N(0) Z(0) C(1) A(0x01)\r
+H(1) C(0) A(0xA1) -> V(0) N(0) Z(0) C(1) A(0x07)\r
+H(1) C(1) A(0xA1) -> V(0) N(0) Z(0) C(1) A(0x07)\r
+\r
+H(0) C(0) A(0xA2) -> V(0) N(0) Z(0) C(1) A(0x02)\r
+H(0) C(1) A(0xA2) -> V(0) N(0) Z(0) C(1) A(0x02)\r
+H(1) C(0) A(0xA2) -> V(0) N(0) Z(0) C(1) A(0x08)\r
+H(1) C(1) A(0xA2) -> V(0) N(0) Z(0) C(1) A(0x08)\r
+\r
+H(0) C(0) A(0xA3) -> V(0) N(0) Z(0) C(1) A(0x03)\r
+H(0) C(1) A(0xA3) -> V(0) N(0) Z(0) C(1) A(0x03)\r
+H(1) C(0) A(0xA3) -> V(0) N(0) Z(0) C(1) A(0x09)\r
+H(1) C(1) A(0xA3) -> V(0) N(0) Z(0) C(1) A(0x09)\r
+\r
+H(0) C(0) A(0xA4) -> V(0) N(0) Z(0) C(1) A(0x04)\r
+H(0) C(1) A(0xA4) -> V(0) N(0) Z(0) C(1) A(0x04)\r
+H(1) C(0) A(0xA4) -> V(0) N(0) Z(0) C(1) A(0x0A)\r
+H(1) C(1) A(0xA4) -> V(0) N(0) Z(0) C(1) A(0x0A)\r
+\r
+H(0) C(0) A(0xA5) -> V(0) N(0) Z(0) C(1) A(0x05)\r
+H(0) C(1) A(0xA5) -> V(0) N(0) Z(0) C(1) A(0x05)\r
+H(1) C(0) A(0xA5) -> V(0) N(0) Z(0) C(1) A(0x0B)\r
+H(1) C(1) A(0xA5) -> V(0) N(0) Z(0) C(1) A(0x0B)\r
+\r
+H(0) C(0) A(0xA6) -> V(0) N(0) Z(0) C(1) A(0x06)\r
+H(0) C(1) A(0xA6) -> V(0) N(0) Z(0) C(1) A(0x06)\r
+H(1) C(0) A(0xA6) -> V(0) N(0) Z(0) C(1) A(0x0C)\r
+H(1) C(1) A(0xA6) -> V(0) N(0) Z(0) C(1) A(0x0C)\r
+\r
+H(0) C(0) A(0xA7) -> V(0) N(0) Z(0) C(1) A(0x07)\r
+H(0) C(1) A(0xA7) -> V(0) N(0) Z(0) C(1) A(0x07)\r
+H(1) C(0) A(0xA7) -> V(0) N(0) Z(0) C(1) A(0x0D)\r
+H(1) C(1) A(0xA7) -> V(0) N(0) Z(0) C(1) A(0x0D)\r
+\r
+H(0) C(0) A(0xA8) -> V(0) N(0) Z(0) C(1) A(0x08)\r
+H(0) C(1) A(0xA8) -> V(0) N(0) Z(0) C(1) A(0x08)\r
+H(1) C(0) A(0xA8) -> V(0) N(0) Z(0) C(1) A(0x0E)\r
+H(1) C(1) A(0xA8) -> V(0) N(0) Z(0) C(1) A(0x0E)\r
+\r
+H(0) C(0) A(0xA9) -> V(0) N(0) Z(0) C(1) A(0x09)\r
+H(0) C(1) A(0xA9) -> V(0) N(0) Z(0) C(1) A(0x09)\r
+H(1) C(0) A(0xA9) -> V(0) N(0) Z(0) C(1) A(0x0F)\r
+H(1) C(1) A(0xA9) -> V(0) N(0) Z(0) C(1) A(0x0F)\r
+\r
+H(0) C(0) A(0xAA) -> V(0) N(0) Z(0) C(1) A(0x10)\r
+H(0) C(1) A(0xAA) -> V(0) N(0) Z(0) C(1) A(0x10)\r
+H(1) C(0) A(0xAA) -> V(0) N(0) Z(0) C(1) A(0x10)\r
+H(1) C(1) A(0xAA) -> V(0) N(0) Z(0) C(1) A(0x10)\r
+\r
+H(0) C(0) A(0xAB) -> V(0) N(0) Z(0) C(1) A(0x11)\r
+H(0) C(1) A(0xAB) -> V(0) N(0) Z(0) C(1) A(0x11)\r
+H(1) C(0) A(0xAB) -> V(0) N(0) Z(0) C(1) A(0x11)\r
+H(1) C(1) A(0xAB) -> V(0) N(0) Z(0) C(1) A(0x11)\r
+\r
+H(0) C(0) A(0xAC) -> V(0) N(0) Z(0) C(1) A(0x12)\r
+H(0) C(1) A(0xAC) -> V(0) N(0) Z(0) C(1) A(0x12)\r
+H(1) C(0) A(0xAC) -> V(0) N(0) Z(0) C(1) A(0x12)\r
+H(1) C(1) A(0xAC) -> V(0) N(0) Z(0) C(1) A(0x12)\r
+\r
+H(0) C(0) A(0xAD) -> V(0) N(0) Z(0) C(1) A(0x13)\r
+H(0) C(1) A(0xAD) -> V(0) N(0) Z(0) C(1) A(0x13)\r
+H(1) C(0) A(0xAD) -> V(0) N(0) Z(0) C(1) A(0x13)\r
+H(1) C(1) A(0xAD) -> V(0) N(0) Z(0) C(1) A(0x13)\r
+\r
+H(0) C(0) A(0xAE) -> V(0) N(0) Z(0) C(1) A(0x14)\r
+H(0) C(1) A(0xAE) -> V(0) N(0) Z(0) C(1) A(0x14)\r
+H(1) C(0) A(0xAE) -> V(0) N(0) Z(0) C(1) A(0x14)\r
+H(1) C(1) A(0xAE) -> V(0) N(0) Z(0) C(1) A(0x14)\r
+\r
+H(0) C(0) A(0xAF) -> V(0) N(0) Z(0) C(1) A(0x15)\r
+H(0) C(1) A(0xAF) -> V(0) N(0) Z(0) C(1) A(0x15)\r
+H(1) C(0) A(0xAF) -> V(0) N(0) Z(0) C(1) A(0x15)\r
+H(1) C(1) A(0xAF) -> V(0) N(0) Z(0) C(1) A(0x15)\r
+\r
+H(0) C(0) A(0xB0) -> V(0) N(0) Z(0) C(1) A(0x10)\r
+H(0) C(1) A(0xB0) -> V(0) N(0) Z(0) C(1) A(0x10)\r
+H(1) C(0) A(0xB0) -> V(0) N(0) Z(0) C(1) A(0x16)\r
+H(1) C(1) A(0xB0) -> V(0) N(0) Z(0) C(1) A(0x16)\r
+\r
+H(0) C(0) A(0xB1) -> V(0) N(0) Z(0) C(1) A(0x11)\r
+H(0) C(1) A(0xB1) -> V(0) N(0) Z(0) C(1) A(0x11)\r
+H(1) C(0) A(0xB1) -> V(0) N(0) Z(0) C(1) A(0x17)\r
+H(1) C(1) A(0xB1) -> V(0) N(0) Z(0) C(1) A(0x17)\r
+\r
+H(0) C(0) A(0xB2) -> V(0) N(0) Z(0) C(1) A(0x12)\r
+H(0) C(1) A(0xB2) -> V(0) N(0) Z(0) C(1) A(0x12)\r
+H(1) C(0) A(0xB2) -> V(0) N(0) Z(0) C(1) A(0x18)\r
+H(1) C(1) A(0xB2) -> V(0) N(0) Z(0) C(1) A(0x18)\r
+\r
+H(0) C(0) A(0xB3) -> V(0) N(0) Z(0) C(1) A(0x13)\r
+H(0) C(1) A(0xB3) -> V(0) N(0) Z(0) C(1) A(0x13)\r
+H(1) C(0) A(0xB3) -> V(0) N(0) Z(0) C(1) A(0x19)\r
+H(1) C(1) A(0xB3) -> V(0) N(0) Z(0) C(1) A(0x19)\r
+\r
+H(0) C(0) A(0xB4) -> V(0) N(0) Z(0) C(1) A(0x14)\r
+H(0) C(1) A(0xB4) -> V(0) N(0) Z(0) C(1) A(0x14)\r
+H(1) C(0) A(0xB4) -> V(0) N(0) Z(0) C(1) A(0x1A)\r
+H(1) C(1) A(0xB4) -> V(0) N(0) Z(0) C(1) A(0x1A)\r
+\r
+H(0) C(0) A(0xB5) -> V(0) N(0) Z(0) C(1) A(0x15)\r
+H(0) C(1) A(0xB5) -> V(0) N(0) Z(0) C(1) A(0x15)\r
+H(1) C(0) A(0xB5) -> V(0) N(0) Z(0) C(1) A(0x1B)\r
+H(1) C(1) A(0xB5) -> V(0) N(0) Z(0) C(1) A(0x1B)\r
+\r
+H(0) C(0) A(0xB6) -> V(0) N(0) Z(0) C(1) A(0x16)\r
+H(0) C(1) A(0xB6) -> V(0) N(0) Z(0) C(1) A(0x16)\r
+H(1) C(0) A(0xB6) -> V(0) N(0) Z(0) C(1) A(0x1C)\r
+H(1) C(1) A(0xB6) -> V(0) N(0) Z(0) C(1) A(0x1C)\r
+\r
+H(0) C(0) A(0xB7) -> V(0) N(0) Z(0) C(1) A(0x17)\r
+H(0) C(1) A(0xB7) -> V(0) N(0) Z(0) C(1) A(0x17)\r
+H(1) C(0) A(0xB7) -> V(0) N(0) Z(0) C(1) A(0x1D)\r
+H(1) C(1) A(0xB7) -> V(0) N(0) Z(0) C(1) A(0x1D)\r
+\r
+H(0) C(0) A(0xB8) -> V(0) N(0) Z(0) C(1) A(0x18)\r
+H(0) C(1) A(0xB8) -> V(0) N(0) Z(0) C(1) A(0x18)\r
+H(1) C(0) A(0xB8) -> V(0) N(0) Z(0) C(1) A(0x1E)\r
+H(1) C(1) A(0xB8) -> V(0) N(0) Z(0) C(1) A(0x1E)\r
+\r
+H(0) C(0) A(0xB9) -> V(0) N(0) Z(0) C(1) A(0x19)\r
+H(0) C(1) A(0xB9) -> V(0) N(0) Z(0) C(1) A(0x19)\r
+H(1) C(0) A(0xB9) -> V(0) N(0) Z(0) C(1) A(0x1F)\r
+H(1) C(1) A(0xB9) -> V(0) N(0) Z(0) C(1) A(0x1F)\r
+\r
+H(0) C(0) A(0xBA) -> V(0) N(0) Z(0) C(1) A(0x20)\r
+H(0) C(1) A(0xBA) -> V(0) N(0) Z(0) C(1) A(0x20)\r
+H(1) C(0) A(0xBA) -> V(0) N(0) Z(0) C(1) A(0x20)\r
+H(1) C(1) A(0xBA) -> V(0) N(0) Z(0) C(1) A(0x20)\r
+\r
+H(0) C(0) A(0xBB) -> V(0) N(0) Z(0) C(1) A(0x21)\r
+H(0) C(1) A(0xBB) -> V(0) N(0) Z(0) C(1) A(0x21)\r
+H(1) C(0) A(0xBB) -> V(0) N(0) Z(0) C(1) A(0x21)\r
+H(1) C(1) A(0xBB) -> V(0) N(0) Z(0) C(1) A(0x21)\r
+\r
+H(0) C(0) A(0xBC) -> V(0) N(0) Z(0) C(1) A(0x22)\r
+H(0) C(1) A(0xBC) -> V(0) N(0) Z(0) C(1) A(0x22)\r
+H(1) C(0) A(0xBC) -> V(0) N(0) Z(0) C(1) A(0x22)\r
+H(1) C(1) A(0xBC) -> V(0) N(0) Z(0) C(1) A(0x22)\r
+\r
+H(0) C(0) A(0xBD) -> V(0) N(0) Z(0) C(1) A(0x23)\r
+H(0) C(1) A(0xBD) -> V(0) N(0) Z(0) C(1) A(0x23)\r
+H(1) C(0) A(0xBD) -> V(0) N(0) Z(0) C(1) A(0x23)\r
+H(1) C(1) A(0xBD) -> V(0) N(0) Z(0) C(1) A(0x23)\r
+\r
+H(0) C(0) A(0xBE) -> V(0) N(0) Z(0) C(1) A(0x24)\r
+H(0) C(1) A(0xBE) -> V(0) N(0) Z(0) C(1) A(0x24)\r
+H(1) C(0) A(0xBE) -> V(0) N(0) Z(0) C(1) A(0x24)\r
+H(1) C(1) A(0xBE) -> V(0) N(0) Z(0) C(1) A(0x24)\r
+\r
+H(0) C(0) A(0xBF) -> V(0) N(0) Z(0) C(1) A(0x25)\r
+H(0) C(1) A(0xBF) -> V(0) N(0) Z(0) C(1) A(0x25)\r
+H(1) C(0) A(0xBF) -> V(0) N(0) Z(0) C(1) A(0x25)\r
+H(1) C(1) A(0xBF) -> V(0) N(0) Z(0) C(1) A(0x25)\r
+\r
+H(0) C(0) A(0xC0) -> V(0) N(0) Z(0) C(1) A(0x20)\r
+H(0) C(1) A(0xC0) -> V(0) N(0) Z(0) C(1) A(0x20)\r
+H(1) C(0) A(0xC0) -> V(0) N(0) Z(0) C(1) A(0x26)\r
+H(1) C(1) A(0xC0) -> V(0) N(0) Z(0) C(1) A(0x26)\r
+\r
+H(0) C(0) A(0xC1) -> V(0) N(0) Z(0) C(1) A(0x21)\r
+H(0) C(1) A(0xC1) -> V(0) N(0) Z(0) C(1) A(0x21)\r
+H(1) C(0) A(0xC1) -> V(0) N(0) Z(0) C(1) A(0x27)\r
+H(1) C(1) A(0xC1) -> V(0) N(0) Z(0) C(1) A(0x27)\r
+\r
+H(0) C(0) A(0xC2) -> V(0) N(0) Z(0) C(1) A(0x22)\r
+H(0) C(1) A(0xC2) -> V(0) N(0) Z(0) C(1) A(0x22)\r
+H(1) C(0) A(0xC2) -> V(0) N(0) Z(0) C(1) A(0x28)\r
+H(1) C(1) A(0xC2) -> V(0) N(0) Z(0) C(1) A(0x28)\r
+\r
+H(0) C(0) A(0xC3) -> V(0) N(0) Z(0) C(1) A(0x23)\r
+H(0) C(1) A(0xC3) -> V(0) N(0) Z(0) C(1) A(0x23)\r
+H(1) C(0) A(0xC3) -> V(0) N(0) Z(0) C(1) A(0x29)\r
+H(1) C(1) A(0xC3) -> V(0) N(0) Z(0) C(1) A(0x29)\r
+\r
+H(0) C(0) A(0xC4) -> V(0) N(0) Z(0) C(1) A(0x24)\r
+H(0) C(1) A(0xC4) -> V(0) N(0) Z(0) C(1) A(0x24)\r
+H(1) C(0) A(0xC4) -> V(0) N(0) Z(0) C(1) A(0x2A)\r
+H(1) C(1) A(0xC4) -> V(0) N(0) Z(0) C(1) A(0x2A)\r
+\r
+H(0) C(0) A(0xC5) -> V(0) N(0) Z(0) C(1) A(0x25)\r
+H(0) C(1) A(0xC5) -> V(0) N(0) Z(0) C(1) A(0x25)\r
+H(1) C(0) A(0xC5) -> V(0) N(0) Z(0) C(1) A(0x2B)\r
+H(1) C(1) A(0xC5) -> V(0) N(0) Z(0) C(1) A(0x2B)\r
+\r
+H(0) C(0) A(0xC6) -> V(0) N(0) Z(0) C(1) A(0x26)\r
+H(0) C(1) A(0xC6) -> V(0) N(0) Z(0) C(1) A(0x26)\r
+H(1) C(0) A(0xC6) -> V(0) N(0) Z(0) C(1) A(0x2C)\r
+H(1) C(1) A(0xC6) -> V(0) N(0) Z(0) C(1) A(0x2C)\r
+\r
+H(0) C(0) A(0xC7) -> V(0) N(0) Z(0) C(1) A(0x27)\r
+H(0) C(1) A(0xC7) -> V(0) N(0) Z(0) C(1) A(0x27)\r
+H(1) C(0) A(0xC7) -> V(0) N(0) Z(0) C(1) A(0x2D)\r
+H(1) C(1) A(0xC7) -> V(0) N(0) Z(0) C(1) A(0x2D)\r
+\r
+H(0) C(0) A(0xC8) -> V(0) N(0) Z(0) C(1) A(0x28)\r
+H(0) C(1) A(0xC8) -> V(0) N(0) Z(0) C(1) A(0x28)\r
+H(1) C(0) A(0xC8) -> V(0) N(0) Z(0) C(1) A(0x2E)\r
+H(1) C(1) A(0xC8) -> V(0) N(0) Z(0) C(1) A(0x2E)\r
+\r
+H(0) C(0) A(0xC9) -> V(0) N(0) Z(0) C(1) A(0x29)\r
+H(0) C(1) A(0xC9) -> V(0) N(0) Z(0) C(1) A(0x29)\r
+H(1) C(0) A(0xC9) -> V(0) N(0) Z(0) C(1) A(0x2F)\r
+H(1) C(1) A(0xC9) -> V(0) N(0) Z(0) C(1) A(0x2F)\r
+\r
+H(0) C(0) A(0xCA) -> V(0) N(0) Z(0) C(1) A(0x30)\r
+H(0) C(1) A(0xCA) -> V(0) N(0) Z(0) C(1) A(0x30)\r
+H(1) C(0) A(0xCA) -> V(0) N(0) Z(0) C(1) A(0x30)\r
+H(1) C(1) A(0xCA) -> V(0) N(0) Z(0) C(1) A(0x30)\r
+\r
+H(0) C(0) A(0xCB) -> V(0) N(0) Z(0) C(1) A(0x31)\r
+H(0) C(1) A(0xCB) -> V(0) N(0) Z(0) C(1) A(0x31)\r
+H(1) C(0) A(0xCB) -> V(0) N(0) Z(0) C(1) A(0x31)\r
+H(1) C(1) A(0xCB) -> V(0) N(0) Z(0) C(1) A(0x31)\r
+\r
+H(0) C(0) A(0xCC) -> V(0) N(0) Z(0) C(1) A(0x32)\r
+H(0) C(1) A(0xCC) -> V(0) N(0) Z(0) C(1) A(0x32)\r
+H(1) C(0) A(0xCC) -> V(0) N(0) Z(0) C(1) A(0x32)\r
+H(1) C(1) A(0xCC) -> V(0) N(0) Z(0) C(1) A(0x32)\r
+\r
+H(0) C(0) A(0xCD) -> V(0) N(0) Z(0) C(1) A(0x33)\r
+H(0) C(1) A(0xCD) -> V(0) N(0) Z(0) C(1) A(0x33)\r
+H(1) C(0) A(0xCD) -> V(0) N(0) Z(0) C(1) A(0x33)\r
+H(1) C(1) A(0xCD) -> V(0) N(0) Z(0) C(1) A(0x33)\r
+\r
+H(0) C(0) A(0xCE) -> V(0) N(0) Z(0) C(1) A(0x34)\r
+H(0) C(1) A(0xCE) -> V(0) N(0) Z(0) C(1) A(0x34)\r
+H(1) C(0) A(0xCE) -> V(0) N(0) Z(0) C(1) A(0x34)\r
+H(1) C(1) A(0xCE) -> V(0) N(0) Z(0) C(1) A(0x34)\r
+\r
+H(0) C(0) A(0xCF) -> V(0) N(0) Z(0) C(1) A(0x35)\r
+H(0) C(1) A(0xCF) -> V(0) N(0) Z(0) C(1) A(0x35)\r
+H(1) C(0) A(0xCF) -> V(0) N(0) Z(0) C(1) A(0x35)\r
+H(1) C(1) A(0xCF) -> V(0) N(0) Z(0) C(1) A(0x35)\r
+\r
+H(0) C(0) A(0xD0) -> V(0) N(0) Z(0) C(1) A(0x30)\r
+H(0) C(1) A(0xD0) -> V(0) N(0) Z(0) C(1) A(0x30)\r
+H(1) C(0) A(0xD0) -> V(0) N(0) Z(0) C(1) A(0x36)\r
+H(1) C(1) A(0xD0) -> V(0) N(0) Z(0) C(1) A(0x36)\r
+\r
+H(0) C(0) A(0xD1) -> V(0) N(0) Z(0) C(1) A(0x31)\r
+H(0) C(1) A(0xD1) -> V(0) N(0) Z(0) C(1) A(0x31)\r
+H(1) C(0) A(0xD1) -> V(0) N(0) Z(0) C(1) A(0x37)\r
+H(1) C(1) A(0xD1) -> V(0) N(0) Z(0) C(1) A(0x37)\r
+\r
+H(0) C(0) A(0xD2) -> V(0) N(0) Z(0) C(1) A(0x32)\r
+H(0) C(1) A(0xD2) -> V(0) N(0) Z(0) C(1) A(0x32)\r
+H(1) C(0) A(0xD2) -> V(0) N(0) Z(0) C(1) A(0x38)\r
+H(1) C(1) A(0xD2) -> V(0) N(0) Z(0) C(1) A(0x38)\r
+\r
+H(0) C(0) A(0xD3) -> V(0) N(0) Z(0) C(1) A(0x33)\r
+H(0) C(1) A(0xD3) -> V(0) N(0) Z(0) C(1) A(0x33)\r
+H(1) C(0) A(0xD3) -> V(0) N(0) Z(0) C(1) A(0x39)\r
+H(1) C(1) A(0xD3) -> V(0) N(0) Z(0) C(1) A(0x39)\r
+\r
+H(0) C(0) A(0xD4) -> V(0) N(0) Z(0) C(1) A(0x34)\r
+H(0) C(1) A(0xD4) -> V(0) N(0) Z(0) C(1) A(0x34)\r
+H(1) C(0) A(0xD4) -> V(0) N(0) Z(0) C(1) A(0x3A)\r
+H(1) C(1) A(0xD4) -> V(0) N(0) Z(0) C(1) A(0x3A)\r
+\r
+H(0) C(0) A(0xD5) -> V(0) N(0) Z(0) C(1) A(0x35)\r
+H(0) C(1) A(0xD5) -> V(0) N(0) Z(0) C(1) A(0x35)\r
+H(1) C(0) A(0xD5) -> V(0) N(0) Z(0) C(1) A(0x3B)\r
+H(1) C(1) A(0xD5) -> V(0) N(0) Z(0) C(1) A(0x3B)\r
+\r
+H(0) C(0) A(0xD6) -> V(0) N(0) Z(0) C(1) A(0x36)\r
+H(0) C(1) A(0xD6) -> V(0) N(0) Z(0) C(1) A(0x36)\r
+H(1) C(0) A(0xD6) -> V(0) N(0) Z(0) C(1) A(0x3C)\r
+H(1) C(1) A(0xD6) -> V(0) N(0) Z(0) C(1) A(0x3C)\r
+\r
+H(0) C(0) A(0xD7) -> V(0) N(0) Z(0) C(1) A(0x37)\r
+H(0) C(1) A(0xD7) -> V(0) N(0) Z(0) C(1) A(0x37)\r
+H(1) C(0) A(0xD7) -> V(0) N(0) Z(0) C(1) A(0x3D)\r
+H(1) C(1) A(0xD7) -> V(0) N(0) Z(0) C(1) A(0x3D)\r
+\r
+H(0) C(0) A(0xD8) -> V(0) N(0) Z(0) C(1) A(0x38)\r
+H(0) C(1) A(0xD8) -> V(0) N(0) Z(0) C(1) A(0x38)\r
+H(1) C(0) A(0xD8) -> V(0) N(0) Z(0) C(1) A(0x3E)\r
+H(1) C(1) A(0xD8) -> V(0) N(0) Z(0) C(1) A(0x3E)\r
+\r
+H(0) C(0) A(0xD9) -> V(0) N(0) Z(0) C(1) A(0x39)\r
+H(0) C(1) A(0xD9) -> V(0) N(0) Z(0) C(1) A(0x39)\r
+H(1) C(0) A(0xD9) -> V(0) N(0) Z(0) C(1) A(0x3F)\r
+H(1) C(1) A(0xD9) -> V(0) N(0) Z(0) C(1) A(0x3F)\r
+\r
+H(0) C(0) A(0xDA) -> V(0) N(0) Z(0) C(1) A(0x40)\r
+H(0) C(1) A(0xDA) -> V(0) N(0) Z(0) C(1) A(0x40)\r
+H(1) C(0) A(0xDA) -> V(0) N(0) Z(0) C(1) A(0x40)\r
+H(1) C(1) A(0xDA) -> V(0) N(0) Z(0) C(1) A(0x40)\r
+\r
+H(0) C(0) A(0xDB) -> V(0) N(0) Z(0) C(1) A(0x41)\r
+H(0) C(1) A(0xDB) -> V(0) N(0) Z(0) C(1) A(0x41)\r
+H(1) C(0) A(0xDB) -> V(0) N(0) Z(0) C(1) A(0x41)\r
+H(1) C(1) A(0xDB) -> V(0) N(0) Z(0) C(1) A(0x41)\r
+\r
+H(0) C(0) A(0xDC) -> V(0) N(0) Z(0) C(1) A(0x42)\r
+H(0) C(1) A(0xDC) -> V(0) N(0) Z(0) C(1) A(0x42)\r
+H(1) C(0) A(0xDC) -> V(0) N(0) Z(0) C(1) A(0x42)\r
+H(1) C(1) A(0xDC) -> V(0) N(0) Z(0) C(1) A(0x42)\r
+\r
+H(0) C(0) A(0xDD) -> V(0) N(0) Z(0) C(1) A(0x43)\r
+H(0) C(1) A(0xDD) -> V(0) N(0) Z(0) C(1) A(0x43)\r
+H(1) C(0) A(0xDD) -> V(0) N(0) Z(0) C(1) A(0x43)\r
+H(1) C(1) A(0xDD) -> V(0) N(0) Z(0) C(1) A(0x43)\r
+\r
+H(0) C(0) A(0xDE) -> V(0) N(0) Z(0) C(1) A(0x44)\r
+H(0) C(1) A(0xDE) -> V(0) N(0) Z(0) C(1) A(0x44)\r
+H(1) C(0) A(0xDE) -> V(0) N(0) Z(0) C(1) A(0x44)\r
+H(1) C(1) A(0xDE) -> V(0) N(0) Z(0) C(1) A(0x44)\r
+\r
+H(0) C(0) A(0xDF) -> V(0) N(0) Z(0) C(1) A(0x45)\r
+H(0) C(1) A(0xDF) -> V(0) N(0) Z(0) C(1) A(0x45)\r
+H(1) C(0) A(0xDF) -> V(0) N(0) Z(0) C(1) A(0x45)\r
+H(1) C(1) A(0xDF) -> V(0) N(0) Z(0) C(1) A(0x45)\r
+\r
+H(0) C(0) A(0xE0) -> V(0) N(0) Z(0) C(1) A(0x40)\r
+H(0) C(1) A(0xE0) -> V(0) N(0) Z(0) C(1) A(0x40)\r
+H(1) C(0) A(0xE0) -> V(0) N(0) Z(0) C(1) A(0x46)\r
+H(1) C(1) A(0xE0) -> V(0) N(0) Z(0) C(1) A(0x46)\r
+\r
+H(0) C(0) A(0xE1) -> V(0) N(0) Z(0) C(1) A(0x41)\r
+H(0) C(1) A(0xE1) -> V(0) N(0) Z(0) C(1) A(0x41)\r
+H(1) C(0) A(0xE1) -> V(0) N(0) Z(0) C(1) A(0x47)\r
+H(1) C(1) A(0xE1) -> V(0) N(0) Z(0) C(1) A(0x47)\r
+\r
+H(0) C(0) A(0xE2) -> V(0) N(0) Z(0) C(1) A(0x42)\r
+H(0) C(1) A(0xE2) -> V(0) N(0) Z(0) C(1) A(0x42)\r
+H(1) C(0) A(0xE2) -> V(0) N(0) Z(0) C(1) A(0x48)\r
+H(1) C(1) A(0xE2) -> V(0) N(0) Z(0) C(1) A(0x48)\r
+\r
+H(0) C(0) A(0xE3) -> V(0) N(0) Z(0) C(1) A(0x43)\r
+H(0) C(1) A(0xE3) -> V(0) N(0) Z(0) C(1) A(0x43)\r
+H(1) C(0) A(0xE3) -> V(0) N(0) Z(0) C(1) A(0x49)\r
+H(1) C(1) A(0xE3) -> V(0) N(0) Z(0) C(1) A(0x49)\r
+\r
+H(0) C(0) A(0xE4) -> V(0) N(0) Z(0) C(1) A(0x44)\r
+H(0) C(1) A(0xE4) -> V(0) N(0) Z(0) C(1) A(0x44)\r
+H(1) C(0) A(0xE4) -> V(0) N(0) Z(0) C(1) A(0x4A)\r
+H(1) C(1) A(0xE4) -> V(0) N(0) Z(0) C(1) A(0x4A)\r
+\r
+H(0) C(0) A(0xE5) -> V(0) N(0) Z(0) C(1) A(0x45)\r
+H(0) C(1) A(0xE5) -> V(0) N(0) Z(0) C(1) A(0x45)\r
+H(1) C(0) A(0xE5) -> V(0) N(0) Z(0) C(1) A(0x4B)\r
+H(1) C(1) A(0xE5) -> V(0) N(0) Z(0) C(1) A(0x4B)\r
+\r
+H(0) C(0) A(0xE6) -> V(0) N(0) Z(0) C(1) A(0x46)\r
+H(0) C(1) A(0xE6) -> V(0) N(0) Z(0) C(1) A(0x46)\r
+H(1) C(0) A(0xE6) -> V(0) N(0) Z(0) C(1) A(0x4C)\r
+H(1) C(1) A(0xE6) -> V(0) N(0) Z(0) C(1) A(0x4C)\r
+\r
+H(0) C(0) A(0xE7) -> V(0) N(0) Z(0) C(1) A(0x47)\r
+H(0) C(1) A(0xE7) -> V(0) N(0) Z(0) C(1) A(0x47)\r
+H(1) C(0) A(0xE7) -> V(0) N(0) Z(0) C(1) A(0x4D)\r
+H(1) C(1) A(0xE7) -> V(0) N(0) Z(0) C(1) A(0x4D)\r
+\r
+H(0) C(0) A(0xE8) -> V(0) N(0) Z(0) C(1) A(0x48)\r
+H(0) C(1) A(0xE8) -> V(0) N(0) Z(0) C(1) A(0x48)\r
+H(1) C(0) A(0xE8) -> V(0) N(0) Z(0) C(1) A(0x4E)\r
+H(1) C(1) A(0xE8) -> V(0) N(0) Z(0) C(1) A(0x4E)\r
+\r
+H(0) C(0) A(0xE9) -> V(0) N(0) Z(0) C(1) A(0x49)\r
+H(0) C(1) A(0xE9) -> V(0) N(0) Z(0) C(1) A(0x49)\r
+H(1) C(0) A(0xE9) -> V(0) N(0) Z(0) C(1) A(0x4F)\r
+H(1) C(1) A(0xE9) -> V(0) N(0) Z(0) C(1) A(0x4F)\r
+\r
+H(0) C(0) A(0xEA) -> V(0) N(0) Z(0) C(1) A(0x50)\r
+H(0) C(1) A(0xEA) -> V(0) N(0) Z(0) C(1) A(0x50)\r
+H(1) C(0) A(0xEA) -> V(0) N(0) Z(0) C(1) A(0x50)\r
+H(1) C(1) A(0xEA) -> V(0) N(0) Z(0) C(1) A(0x50)\r
+\r
+H(0) C(0) A(0xEB) -> V(0) N(0) Z(0) C(1) A(0x51)\r
+H(0) C(1) A(0xEB) -> V(0) N(0) Z(0) C(1) A(0x51)\r
+H(1) C(0) A(0xEB) -> V(0) N(0) Z(0) C(1) A(0x51)\r
+H(1) C(1) A(0xEB) -> V(0) N(0) Z(0) C(1) A(0x51)\r
+\r
+H(0) C(0) A(0xEC) -> V(0) N(0) Z(0) C(1) A(0x52)\r
+H(0) C(1) A(0xEC) -> V(0) N(0) Z(0) C(1) A(0x52)\r
+H(1) C(0) A(0xEC) -> V(0) N(0) Z(0) C(1) A(0x52)\r
+H(1) C(1) A(0xEC) -> V(0) N(0) Z(0) C(1) A(0x52)\r
+\r
+H(0) C(0) A(0xED) -> V(0) N(0) Z(0) C(1) A(0x53)\r
+H(0) C(1) A(0xED) -> V(0) N(0) Z(0) C(1) A(0x53)\r
+H(1) C(0) A(0xED) -> V(0) N(0) Z(0) C(1) A(0x53)\r
+H(1) C(1) A(0xED) -> V(0) N(0) Z(0) C(1) A(0x53)\r
+\r
+H(0) C(0) A(0xEE) -> V(0) N(0) Z(0) C(1) A(0x54)\r
+H(0) C(1) A(0xEE) -> V(0) N(0) Z(0) C(1) A(0x54)\r
+H(1) C(0) A(0xEE) -> V(0) N(0) Z(0) C(1) A(0x54)\r
+H(1) C(1) A(0xEE) -> V(0) N(0) Z(0) C(1) A(0x54)\r
+\r
+H(0) C(0) A(0xEF) -> V(0) N(0) Z(0) C(1) A(0x55)\r
+H(0) C(1) A(0xEF) -> V(0) N(0) Z(0) C(1) A(0x55)\r
+H(1) C(0) A(0xEF) -> V(0) N(0) Z(0) C(1) A(0x55)\r
+H(1) C(1) A(0xEF) -> V(0) N(0) Z(0) C(1) A(0x55)\r
+\r
+H(0) C(0) A(0xF0) -> V(0) N(0) Z(0) C(1) A(0x50)\r
+H(0) C(1) A(0xF0) -> V(0) N(0) Z(0) C(1) A(0x50)\r
+H(1) C(0) A(0xF0) -> V(0) N(0) Z(0) C(1) A(0x56)\r
+H(1) C(1) A(0xF0) -> V(0) N(0) Z(0) C(1) A(0x56)\r
+\r
+H(0) C(0) A(0xF1) -> V(0) N(0) Z(0) C(1) A(0x51)\r
+H(0) C(1) A(0xF1) -> V(0) N(0) Z(0) C(1) A(0x51)\r
+H(1) C(0) A(0xF1) -> V(0) N(0) Z(0) C(1) A(0x57)\r
+H(1) C(1) A(0xF1) -> V(0) N(0) Z(0) C(1) A(0x57)\r
+\r
+H(0) C(0) A(0xF2) -> V(0) N(0) Z(0) C(1) A(0x52)\r
+H(0) C(1) A(0xF2) -> V(0) N(0) Z(0) C(1) A(0x52)\r
+H(1) C(0) A(0xF2) -> V(0) N(0) Z(0) C(1) A(0x58)\r
+H(1) C(1) A(0xF2) -> V(0) N(0) Z(0) C(1) A(0x58)\r
+\r
+H(0) C(0) A(0xF3) -> V(0) N(0) Z(0) C(1) A(0x53)\r
+H(0) C(1) A(0xF3) -> V(0) N(0) Z(0) C(1) A(0x53)\r
+H(1) C(0) A(0xF3) -> V(0) N(0) Z(0) C(1) A(0x59)\r
+H(1) C(1) A(0xF3) -> V(0) N(0) Z(0) C(1) A(0x59)\r
+\r
+H(0) C(0) A(0xF4) -> V(0) N(0) Z(0) C(1) A(0x54)\r
+H(0) C(1) A(0xF4) -> V(0) N(0) Z(0) C(1) A(0x54)\r
+H(1) C(0) A(0xF4) -> V(0) N(0) Z(0) C(1) A(0x5A)\r
+H(1) C(1) A(0xF4) -> V(0) N(0) Z(0) C(1) A(0x5A)\r
+\r
+H(0) C(0) A(0xF5) -> V(0) N(0) Z(0) C(1) A(0x55)\r
+H(0) C(1) A(0xF5) -> V(0) N(0) Z(0) C(1) A(0x55)\r
+H(1) C(0) A(0xF5) -> V(0) N(0) Z(0) C(1) A(0x5B)\r
+H(1) C(1) A(0xF5) -> V(0) N(0) Z(0) C(1) A(0x5B)\r
+\r
+H(0) C(0) A(0xF6) -> V(0) N(0) Z(0) C(1) A(0x56)\r
+H(0) C(1) A(0xF6) -> V(0) N(0) Z(0) C(1) A(0x56)\r
+H(1) C(0) A(0xF6) -> V(0) N(0) Z(0) C(1) A(0x5C)\r
+H(1) C(1) A(0xF6) -> V(0) N(0) Z(0) C(1) A(0x5C)\r
+\r
+H(0) C(0) A(0xF7) -> V(0) N(0) Z(0) C(1) A(0x57)\r
+H(0) C(1) A(0xF7) -> V(0) N(0) Z(0) C(1) A(0x57)\r
+H(1) C(0) A(0xF7) -> V(0) N(0) Z(0) C(1) A(0x5D)\r
+H(1) C(1) A(0xF7) -> V(0) N(0) Z(0) C(1) A(0x5D)\r
+\r
+H(0) C(0) A(0xF8) -> V(0) N(0) Z(0) C(1) A(0x58)\r
+H(0) C(1) A(0xF8) -> V(0) N(0) Z(0) C(1) A(0x58)\r
+H(1) C(0) A(0xF8) -> V(0) N(0) Z(0) C(1) A(0x5E)\r
+H(1) C(1) A(0xF8) -> V(0) N(0) Z(0) C(1) A(0x5E)\r
+\r
+H(0) C(0) A(0xF9) -> V(0) N(0) Z(0) C(1) A(0x59)\r
+H(0) C(1) A(0xF9) -> V(0) N(0) Z(0) C(1) A(0x59)\r
+H(1) C(0) A(0xF9) -> V(0) N(0) Z(0) C(1) A(0x5F)\r
+H(1) C(1) A(0xF9) -> V(0) N(0) Z(0) C(1) A(0x5F)\r
+\r
+H(0) C(0) A(0xFA) -> V(0) N(0) Z(0) C(1) A(0x60)\r
+H(0) C(1) A(0xFA) -> V(0) N(0) Z(0) C(1) A(0x60)\r
+H(1) C(0) A(0xFA) -> V(0) N(0) Z(0) C(1) A(0x60)\r
+H(1) C(1) A(0xFA) -> V(0) N(0) Z(0) C(1) A(0x60)\r
+\r
+H(0) C(0) A(0xFB) -> V(0) N(0) Z(0) C(1) A(0x61)\r
+H(0) C(1) A(0xFB) -> V(0) N(0) Z(0) C(1) A(0x61)\r
+H(1) C(0) A(0xFB) -> V(0) N(0) Z(0) C(1) A(0x61)\r
+H(1) C(1) A(0xFB) -> V(0) N(0) Z(0) C(1) A(0x61)\r
+\r
+H(0) C(0) A(0xFC) -> V(0) N(0) Z(0) C(1) A(0x62)\r
+H(0) C(1) A(0xFC) -> V(0) N(0) Z(0) C(1) A(0x62)\r
+H(1) C(0) A(0xFC) -> V(0) N(0) Z(0) C(1) A(0x62)\r
+H(1) C(1) A(0xFC) -> V(0) N(0) Z(0) C(1) A(0x62)\r
+\r
+H(0) C(0) A(0xFD) -> V(0) N(0) Z(0) C(1) A(0x63)\r
+H(0) C(1) A(0xFD) -> V(0) N(0) Z(0) C(1) A(0x63)\r
+H(1) C(0) A(0xFD) -> V(0) N(0) Z(0) C(1) A(0x63)\r
+H(1) C(1) A(0xFD) -> V(0) N(0) Z(0) C(1) A(0x63)\r
+\r
+H(0) C(0) A(0xFE) -> V(0) N(0) Z(0) C(1) A(0x64)\r
+H(0) C(1) A(0xFE) -> V(0) N(0) Z(0) C(1) A(0x64)\r
+H(1) C(0) A(0xFE) -> V(0) N(0) Z(0) C(1) A(0x64)\r
+H(1) C(1) A(0xFE) -> V(0) N(0) Z(0) C(1) A(0x64)\r
+\r
+H(0) C(0) A(0xFF) -> V(0) N(0) Z(0) C(1) A(0x65)\r
+H(0) C(1) A(0xFF) -> V(0) N(0) Z(0) C(1) A(0x65)\r
+H(1) C(0) A(0xFF) -> V(0) N(0) Z(0) C(1) A(0x65)\r
+H(1) C(1) A(0xFF) -> V(0) N(0) Z(0) C(1) A(0x65)\r
diff --git a/helm/software/matita/library/freescale/doc/freescale.txt b/helm/software/matita/library/freescale/doc/freescale.txt
new file mode 100644 (file)
index 0000000..5a7ea8d
--- /dev/null
@@ -0,0 +1,21 @@
+INDIRIZZI ITALIA\r
+\r
+http://www.freescale.com/webapp/sps/site/overview.jsp?nodeId=0671479183#it\r
+http://www.freescale.com/webapp/sps/site/homepage.jsp?nodeId=06714769568830\r
+\r
+Italy:\r
+Freescale Semiconductor\r
+Via Muzio Attendolo Detto\r
+Sforza, 13\r
+20141 Milano\r
+Italia \r
+\r
+TOOL TEST\r
+\r
+http://www.freescale.com/webapp/sps/site/prod_summary.jsp?code=DEMO9RS08KA2\r
+DEMO9RS08KA2 (50$)\r
+per MC9RS08KA2 (<=1$)\r
+\r
+http://www.freescale.com/webapp/sps/site/prod_summary.jsp?code=USBSPYDER08\r
+USBSPYDER08 (29$)\r
+per MC9RS08KA1/2, MC9S08QD2/4, MC9S08QG4/8 (<=1$)\r
diff --git a/helm/software/matita/library/freescale/doc/ordinamento.txt b/helm/software/matita/library/freescale/doc/ordinamento.txt
new file mode 100644 (file)
index 0000000..cb63d9c
--- /dev/null
@@ -0,0 +1,613 @@
+PROCESSORE\r
+\r
+A : 0x00\r
+PC: 0x18C8 -> 0x1966\r
+SP: 0x0F4B\r
+HX: 0x0F4C\r
+Status: I=1 Z=1\r
+\r
+TEMPI COMPLETI\r
+0   esempi: 183941 -> 219568: 35627 tick\r
+10  esempi: 183941 -> 221074: 37133 =35633+  10*150 tick\r
+20  esempi: 183941 -> 222573: 38632 =35632+  20*150 tick\r
+30  esempi: 183941 -> 224063: 40122 =35622+  30*150 tick\r
+40  esempi: 183941 -> 225553: 41612 =35612+  40*150 tick\r
+50  esempi: 183941 -> 227043: 43102 =35602+  50*150 tick\r
+100 esempi: 183941 -> 234493: 50552 =35552+ 100*150 tick\r
+200 esempi: 183941 -> 249393: 65453 =35453+ 200*150 tick\r
+\r
+oltre bisogna introdurre __RESET_WATCHDOG(); e quindi approssimare ancora di piu'\r
+\r
+500 esempi: 183941 -> 297886: 113945=38945+500 *150 tick\r
+ => con tara di (256+500)*4=3024 e' 35921+...\r
+1024esempi: 183941 -> 378070: 194129=40529+1024*150 tick\r
+ => con tara di (256+1024)*4=5120 e' 35409+...\r
+2048esempi: 183941 -> 534758: 350817=43617+2048*150 tick\r
+ => con tara di (256+2048)*4=9216 e' 34401+...\r
+3072esempi: 183941 -> 691446: 507505=46705+3072*150 tick\r
+ => con tara di (256+3072)*4=13312 e' 33393+...\r
+\r
+si puo' quindi dare un'interpretazione del tipo T(n)=fix+n*esempi=36000+150n\r
+chiaramente utile solo per definire un bound superiore \r
+\r
+NB: per l'ultimo 1s circa di esecuzione con emulatore,\r
+    tutti questi esempi sono stati fatti girare in emulazione\r
+    e non con IN-CIRCUIT debug.\r
+\r
+NB: __RESET_WATCHDOG(); e' in pratica una singola istruzione\r
+    STA 0x1800 cioe' scrivi qualsiasi cosa su locazione di reset (4 tick)\r
+\r
+TEMPI PARZIALI\r
+1) senza ripristino di counters, index, position...\r
+2) calcolati con tara di _RESET_WATCHDOG() sempre per uniformare\r
+\r
+0   : 183941 -> 209557:  25616: ( 25616+   0*4): k          :  25700 >  25616\r
+8   : 183941 -> 210794:  26853: ( 26821+   8*4): k+150.62*n :  26900 >  26821\r
+16  : 183941 -> 212018:  28077: ( 28013+  16*4): k+149.81*n :  28100 >  28013\r
+32  : 183941 -> 214466:  30525: ( 30397+  32*4): k+149.41*n :  30500 >  30397\r
+64  : 183941 -> 219362:  35421: ( 35165+  64*4): k+149.20*n :  35300 >  35165\r
+128 : 183941 -> 229154:  45213: ( 44701+ 128*4): k+149.10*n :  44900 >  44701\r
+256 : 183941 -> 248742:  64801: ( 63777+ 256*4): k+149.07*n :  64100 >  63777\r
+512 : 183941 -> 287914: 103973: (101925+ 512*4): k+149.04*n : 102500 > 101925\r
+1024: 183941 -> 366258: 182317: (178221+1024*4): k+149.03*n : 179300 > 178221\r
+2048: 183941 -> 522946: 339005: (330813+2048*4): k+149.02*n : 332900 > 330813\r
+3072: 183941 -> 679634: 495693: (483405+3072*4): k+149.02*n : 486500 > 483405\r
+\r
+quindi generalizziamo in T(n)=25700+150*n (0x6464+0x96*n)\r
+\r
+ma registri diversi\r
+A : 0x00   -> 0xFF\r
+PC: 0x18C8 -> 0x193B\r
+SP: 0x0F4B\r
+HX: 0x0F4C -> 0x0100\r
+Status: I=1 Z=1\r
+\r
+LOCAZIONI\r
+dati=3072 byte [100-CFF]\r
+counters=256 word [D00-EFF]\r
+index=1 word [F4C-F4D]\r
+position=1 word [F4E-F4F]
+
+(fine RAM a 0x107F)
+=> poiche' counters e stack finali sono difficili da formalizzare meglio azzerare
+   da 0x0D00 a 0x107F, cioe' 380 (1KB - 0x80, la zona iniziale non usata)\r
+\r
+sostanzialmente si puo' riassumere in tesi:\r
+ tutto a 0 tranne dati\r
+\r
+***********************************\r
+\r
+static unsigned int counters[256]={ 0, ... };\r
+static unsigned char dati[3072]={ ... };\r
+\r
+void sort(void)\r
+{\r
+unsigned int index=0,position=0;\r
+\r
+/* prologo: si puo' saltare e inglobare nell'ipotesi */\r
+\r
+for(;index<3072;index++)\r
+    {\r
+    __RESET_WATCHDOG(); /* #solo per esempi > 200 */\r
+    counters[dati[index]]++;\r
+    }\r
+\r
+jmp blocco controllo del for\r
+18C8 20 1D       BRA *+31 ; 18E7\r
+corpo del for\r
+18CA 9E FE 01    LDHX 1,SP\r
+18CD D6 01 00    LDA 256,X\r
+18D0 48          LSLA\r
+18D1 5F          CLRX\r
+18D2 59          ROLX\r
+18D3 AB 00       ADD #0X00\r
+18D5 87          PSHA\r
+18D6 9F          TXA\r
+18D7 A9 0D       ADC #0X0D\r
+18D9 87          PSHA\r
+18DA 8A          PULH\r
+18DB 88          PULX\r
+passo del for\r
+18DC 6C 01       INC 1,X\r
+18DE 26 01       BNE *+3 ; 18E1\r
+18E0 7C          INC ,X\r
+18E1 95          TSX\r
+18E2 6C 01       INC 1,X\r
+18E4 26 01       BNE *+3 ; 18E7\r
+18E6 7C          INC ,X\r
+blocco controllo del for\r
+18E7 9E FE 01    LDHX 1,SP\r
+18EA 65 0C 00    CPHX #0X0C00 ; DIMENSIONE DEI DATI\r
+18ED 25 DB       BCS *-35 ; 18CA\r
+uscita dal for\r
+\r
+for(index=0;index<256;index++)\r
+    {\r
+    __RESET_WATCHDOG(); /* #solo per esempi > 200 */\r
+    while(counters[index]--)\r
+        { dati[position++]=index; }\r
+    }\r
+\r
+init del for\r
+18EF 95          TSX\r
+18F0 6F 01       CLR 1,X\r
+18F2 7F          CLR ,X\r
+salto a condizione del while\r
+18F3 20 0E       BRA *+16 ; 1903\r
+corpo del while\r
+18F5 95          TSX\r
+18F6 E6 01       LDA 1,X\r
+18F8 9E FE 03    LDHX 3,SP\r
+18FB D7 01 00    STA 256,X\r
+18FE AF 01       AIX #1\r
+1900 9E FF 03    STHX 3,SP\r
+condizione del while\r
+1903 95          TSX\r
+1904 EE 01       LDX 1,X\r
+1906 58          LSLX\r
+1907 9E E6 01    LDA 1,SP\r
+190A 49          ROLA\r
+190B 87          PSHA\r
+190C 8A          PULH\r
+190D 89          PSHX\r
+190E 9E BE 0D 00 LDHX 3328,X\r
+1912 89          PSHX\r
+1913 8B          PSHH\r
+1914 AF FF       AIX #-1\r
+1916 8B          PSHH\r
+1917 87          PSHA\r
+1918 8A          PULH\r
+1919 89          PSHX\r
+191A 9E EE 05    LDX 5,SP\r
+191D 86          PULA\r
+191E D7 0D 01    STA 3329,X\r
+1921 86          PULA\r
+1922 D7 0D 00    STA 3328,X\r
+1925 8A          PULH\r
+1926 88          PULX\r
+1927 65 00 00    CPHX #0X0000\r
+192A 8A          PULH\r
+192B 26 C8       BNE *-54 ; 18F5\r
+uscita dal while/passo del for\r
+192D 95          TSX\r
+192E 6C 01       INC 1,X\r
+1930 26 01       BNE *+3 ; 1933\r
+1932 7C          INC ,X\r
+blocco di controllo del for\r
+1933 9E FE 01    LDHX 1,SP\r
+1936 65 01 00    CPHX #0X0100\r
+1939 25 C8       BNE *-54 ; 1903\r
+uscita dal for\r
+\r
+for(index=0;index<256;index++)\r
+    { counters[index]=0; }\r
+\r
+init del for\r
+193B 95          TSX\r
+193C 6F 01       CLR 1,X\r
+193E 7F          CLR ,X\r
+corpo del for\r
+193F 95          TSX\r
+1940 EE 01       LDX 1,X\r
+1942 58          LSLX\r
+1943 9E E6 01    LDA 1,SP\r
+1946 49          ROLA\r
+1947 87          PSHA\r
+1948 8A          PULH\r
+1949 4F          CLRA\r
+194A D7 0D 01    STA 3329,X\r
+194D D7 0D 00    STA 3328,X\r
+passo del for\r
+1950 95          TSX\r
+1951 6C 01       INC 1,X\r
+1953 26 01       BNE *+3 ; 1956\r
+1955 7C          INC ,X\r
+blocco di controllo del for\r
+1956 9E FE 01    LDHX 1,SP\r
+1959 65 01 00    CPHX #0X0100\r
+195C 25 E1       BCS *-29 ; 193F\r
+\r
+position=0;\r
+index=0;\r
+\r
+195E 95          TSX\r
+195F 6F 03       CLR 3,X\r
+1961 6F 02       CLR 2,X\r
+1963 6F 01       CLR 1,X\r
+1965 7F          CLR ,x\r
+\r
+***********************************\r
+\r
+esempio generato con mathematica\r
+\r
+{129, 94, 118, 209, 117, 16, 133, 144, 31, 18, 226, 60, 29, 6, 60, 209, 131, \\r
+235, 114, 31, 144, 240, 71, 163, 208, 70, 197, 50, 201, 11, 27, 115, 17, 175, \\r
+210, 191, 11, 188, 8, 229, 254, 172, 252, 148, 52, 136, 40, 219, 211, 61, 56, \\r
+106, 77, 116, 73, 126, 111, 77, 61, 58, 194, 211, 54, 112, 128, 201, 59, 82, \\r
+143, 30, 132, 82, 45, 175, 27, 1, 53, 114, 16, 7, 93, 171, 224, 118, 189, 60, \\r
+13, 206, 183, 157, 169, 243, 193, 147, 140, 78, 248, 195, 19, 58, 162, 189, \\r
+16, 155, 92, 200, 229, 200, 166, 251, 19, 200, 100, 9, 214, 215, 57, 249, \\r
+228, 3, 73, 206, 113, 87, 52, 46, 163, 83, 249, 42, 240, 225, 102, 137, 46, \\r
+134, 199, 45, 18, 213, 74, 121, 109, 27, 65, 148, 108, 109, 244, 93, 139, \\r
+190, 200, 229, 164, 186, 177, 33, 80, 107, 149, 6, 69, 224, 15, 90, 206, 200, \\r
+130, 79, 194, 245, 248, 19, 218, 43, 121, 191, 164, 90, 210, 118, 141, 31, \\r
+21, 90, 222, 35, 157, 126, 103, 99, 17, 245, 36, 214, 106, 12, 127, 18, 214, \\r
+238, 186, 68, 57, 146, 102, 234, 35, 68, 207, 116, 107, 131, 119, 4, 100, \\r
+179, 60, 38, 252, 213, 65, 183, 36, 224, 180, 19, 50, 69, 24, 219, 0, 182, \\r
+95, 58, 183, 79, 180, 223, 244, 26, 16, 233, 197, 44, 188, 91, 5, 74, 82, \\r
+192, 31, 89, 242, 229, 151, 81, 247, 60, 119, 129, 124, 176, 29, 26, 123, 71, \\r
+121, 194, 63, 211, 182, 215, 201, 146, 13, 127, 11, 117, 14, 81, 159, 76, \\r
+229, 208, 193, 57, 127, 234, 65, 5, 8, 126, 118, 47, 98, 113, 248, 56, 182, \\r
+112, 221, 208, 173, 120, 61, 123, 204, 109, 63, 172, 199, 110, 255, 69, 193, \\r
+116, 191, 41, 73, 197, 124, 45, 242, 204, 250, 88, 172, 88, 81, 14, 72, 112, \\r
+65, 237, 149, 180, 71, 241, 129, 228, 64, 84, 123, 161, 210, 2, 197, 125, \\r
+122, 241, 0, 141, 102, 176, 169, 122, 43, 44, 111, 95, 111, 78, 43, 17, 255, \\r
+71, 127, 207, 211, 133, 184, 77, 153, 131, 118, 232, 48, 248, 208, 89, 184, \\r
+113, 87, 175, 181, 213, 252, 130, 22, 16, 179, 193, 26, 1, 17, 255, 89, 36, \\r
+197, 114, 216, 208, 173, 254, 215, 17, 206, 249, 184, 119, 58, 31, 97, 27, \\r
+187, 85, 128, 124, 37, 52, 137, 242, 201, 223, 53, 197, 18, 240, 5, 222, 38, \\r
+73, 183, 57, 5, 194, 219, 252, 159, 169, 102, 173, 74, 63, 191, 109, 135, \\r
+106, 177, 62, 182, 14, 122, 59, 69, 233, 206, 106, 106, 112, 96, 106, 44, \\r
+210, 184, 54, 33, 0, 84, 49, 96, 27, 76, 202, 190, 95, 129, 183, 155, 38, \\r
+148, 43, 65, 43, 152, 99, 102, 101, 70, 35, 229, 7, 158, 28, 56, 92, 151, 99, \\r
+83, 102, 8, 93, 8, 219, 110, 86, 112, 50, 69, 2, 99, 114, 156, 217, 80, 6, \\r
+89, 23, 104, 210, 215, 142, 105, 95, 129, 132, 135, 220, 152, 229, 181, 195, \\r
+37, 108, 146, 221, 42, 209, 20, 126, 23, 178, 152, 85, 252, 61, 141, 239, \\r
+129, 184, 187, 81, 0, 180, 46, 48, 96, 126, 144, 227, 244, 114, 28, 179, 120, \\r
+27, 159, 27, 3, 163, 5, 222, 56, 186, 14, 237, 236, 31, 56, 227, 247, 170, \\r
+233, 61, 255, 243, 20, 44, 136, 97, 48, 171, 24, 220, 254, 106, 41, 241, 203, \\r
+144, 120, 153, 31, 40, 249, 203, 20, 132, 243, 214, 126, 236, 86, 206, 218, \\r
+94, 97, 241, 102, 105, 147, 89, 60, 29, 107, 244, 89, 77, 56, 169, 59, 127, \\r
+178, 236, 174, 246, 178, 34, 100, 39, 108, 32, 238, 81, 62, 136, 221, 193, \\r
+220, 193, 102, 97, 66, 119, 54, 8, 41, 96, 169, 252, 124, 167, 180, 252, 135, \\r
+29, 108, 162, 63, 29, 23, 15, 2, 38, 162, 106, 92, 237, 39, 197, 123, 245, \\r
+156, 133, 99, 86, 195, 75, 27, 160, 27, 137, 53, 214, 217, 221, 46, 98, 117, \\r
+231, 24, 77, 215, 88, 167, 95, 148, 135, 168, 231, 43, 242, 231, 185, 6, 175, \\r
+218, 220, 198, 63, 141, 121, 149, 218, 91, 146, 238, 60, 254, 73, 90, 16, 77, \\r
+137, 131, 38, 236, 141, 201, 119, 238, 241, 64, 109, 73, 87, 155, 196, 31, \\r
+128, 149, 188, 232, 249, 215, 20, 62, 195, 111, 143, 114, 213, 190, 138, 163, \\r
+247, 96, 58, 116, 251, 189, 116, 131, 227, 157, 217, 184, 19, 80, 64, 138, \\r
+150, 58, 166, 236, 124, 21, 135, 77, 106, 170, 224, 186, 255, 59, 226, 81, \\r
+34, 34, 31, 161, 33, 175, 55, 138, 223, 227, 105, 238, 196, 231, 113, 150, \\r
+17, 228, 57, 229, 175, 245, 87, 235, 85, 101, 139, 62, 141, 70, 83, 178, 25, \\r
+52, 233, 74, 75, 82, 48, 63, 167, 79, 26, 184, 100, 91, 217, 109, 97, 165, \\r
+207, 140, 221, 230, 213, 54, 12, 141, 247, 78, 156, 191, 43, 68, 209, 192, \\r
+128, 8, 173, 206, 189, 76, 83, 101, 182, 72, 246, 100, 124, 152, 16, 157, \\r
+253, 73, 196, 221, 20, 182, 111, 59, 70, 215, 26, 68, 164, 129, 49, 162, 64, \\r
+112, 57, 154, 76, 79, 147, 157, 212, 151, 57, 168, 168, 249, 179, 231, 216, \\r
+77, 109, 136, 107, 84, 85, 146, 25, 228, 177, 227, 70, 213, 108, 249, 51, \\r
+229, 210, 119, 24, 134, 20, 126, 30, 200, 180, 206, 199, 87, 45, 82, 231, 87, \\r
+155, 26, 73, 9, 75, 246, 181, 10, 44, 186, 31, 38, 56, 96, 52, 196, 137, 191, \\r
+4, 95, 244, 17, 243, 74, 226, 216, 54, 250, 67, 12, 96, 81, 45, 3, 252, 146, \\r
+33, 202, 117, 102, 220, 132, 196, 111, 240, 4, 82, 171, 154, 206, 170, 47, \\r
+54, 243, 167, 66, 7, 94, 182, 91, 152, 121, 67, 136, 154, 10, 85, 227, 152, \\r
+117, 162, 234, 185, 50, 65, 7, 56, 177, 3, 78, 111, 151, 60, 84, 187, 188, \\r
+11, 12, 171, 59, 40, 170, 4, 80, 163, 107, 250, 240, 230, 24, 216, 122, 95, \\r
+3, 130, 150, 154, 191, 21, 133, 239, 179, 123, 238, 239, 16, 105, 194, 249, \\r
+45, 129, 243, 14, 195, 127, 221, 46, 42, 189, 234, 146, 255, 250, 123, 211, \\r
+48, 117, 103, 168, 15, 33, 220, 229, 226, 137, 41, 197, 163, 162, 66, 63, \\r
+163, 88, 224, 124, 3, 255, 40, 139, 139, 18, 216, 168, 118, 185, 226, 254, \\r
+33, 63, 172, 70, 188, 248, 211, 235, 252, 159, 231, 97, 203, 191, 78, 196, \\r
+151, 30, 13, 121, 131, 171, 76, 38, 99, 111, 238, 89, 129, 2, 44, 237, 111, \\r
+4, 16, 224, 218, 190, 238, 87, 176, 49, 65, 220, 60, 204, 88, 40, 44, 27, \\r
+134, 214, 249, 213, 74, 234, 11, 32, 44, 66, 206, 69, 43, 1, 170, 177, 110, \\r
+183, 183, 40, 149, 25, 167, 92, 78, 123, 62, 211, 144, 134, 113, 20, 210, \\r
+212, 180, 242, 49, 40, 69, 253, 120, 93, 250, 243, 149, 77, 49, 184, 204, 33, \\r
+25, 66, 42, 242, 186, 4, 159, 67, 69, 28, 116, 187, 112, 94, 1, 172, 109, \\r
+215, 156, 157, 26, 152, 177, 252, 97, 163, 65, 65, 175, 29, 225, 50, 25, 96, \\r
+41, 151, 133, 83, 83, 145, 179, 148, 213, 157, 76, 54, 14, 132, 161, 70, 106, \\r
+31, 243, 107, 190, 74, 25, 117, 252, 198, 234, 126, 209, 51, 103, 183, 238, \\r
+89, 226, 221, 34, 140, 155, 62, 152, 252, 19, 226, 12, 78, 49, 135, 103, 106, \\r
+76, 76, 114, 0, 5, 31, 246, 48, 238, 222, 177, 76, 247, 236, 44, 64, 107, \\r
+104, 144, 136, 111, 179, 71, 98, 146, 155, 43, 50, 64, 167, 137, 64, 35, 92, \\r
+249, 41, 98, 174, 91, 201, 44, 146, 111, 245, 160, 14, 222, 249, 9, 16, 57, \\r
+70, 197, 226, 131, 213, 142, 70, 76, 172, 127, 79, 193, 79, 26, 97, 150, 11, \\r
+0, 111, 38, 140, 238, 147, 27, 158, 165, 150, 42, 235, 70, 95, 60, 214, 45, \\r
+148, 107, 244, 218, 105, 89, 172, 189, 158, 72, 2, 220, 92, 96, 42, 110, 202, \\r
+110, 31, 212, 58, 176, 158, 143, 171, 178, 2, 71, 125, 171, 219, 181, 109, \\r
+226, 137, 77, 4, 190, 250, 34, 20, 113, 18, 27, 13, 186, 90, 108, 30, 58, 15, \\r
+110, 68, 200, 181, 140, 3, 6, 100, 133, 40, 100, 34, 129, 118, 254, 250, 98, \\r
+145, 190, 188, 97, 75, 126, 80, 188, 238, 99, 207, 29, 253, 98, 92, 133, 158, \\r
+165, 38, 231, 70, 59, 234, 190, 4, 136, 242, 146, 11, 217, 233, 41, 56, 136, \\r
+138, 105, 23, 75, 191, 12, 242, 253, 115, 89, 190, 84, 28, 211, 49, 98, 27, \\r
+183, 50, 164, 241, 112, 154, 70, 161, 28, 4, 108, 242, 230, 193, 164, 242, \\r
+42, 75, 53, 155, 153, 255, 1, 19, 249, 92, 60, 81, 138, 165, 95, 158, 94, \\r
+198, 33, 55, 210, 185, 152, 161, 96, 233, 69, 202, 215, 189, 207, 15, 36, \\r
+229, 121, 75, 28, 87, 58, 36, 34, 8, 51, 226, 162, 88, 37, 84, 113, 43, 255, \\r
+237, 72, 246, 35, 49, 186, 81, 159, 170, 204, 3, 21, 199, 45, 211, 235, 143, \\r
+132, 64, 83, 173, 103, 236, 175, 220, 28, 116, 107, 173, 205, 167, 17, 16, \\r
+207, 190, 161, 1, 63, 192, 133, 40, 107, 195, 109, 216, 117, 90, 240, 34, 75, \\r
+156, 49, 228, 231, 198, 252, 48, 213, 249, 26, 64, 31, 109, 213, 120, 181, \\r
+118, 201, 225, 223, 18, 97, 223, 159, 87, 125, 11, 173, 90, 161, 132, 229, \\r
+247, 188, 211, 165, 180, 133, 103, 54, 252, 177, 179, 203, 30, 236, 110, 225, \\r
+28, 165, 83, 152, 246, 223, 65, 19, 46, 127, 14, 56, 60, 212, 140, 42, 40, \\r
+78, 126, 14, 247, 202, 62, 228, 180, 5, 88, 220, 120, 217, 249, 121, 132, 2, \\r
+63, 207, 56, 215, 38, 29, 24, 77, 234, 122, 212, 36, 13, 77, 144, 23, 30, \\r
+110, 182, 204, 192, 176, 94, 153, 109, 207, 238, 41, 207, 162, 12, 171, 116, \\r
+41, 78, 130, 147, 105, 123, 236, 199, 137, 202, 221, 172, 101, 128, 20, 9, \\r
+85, 14, 132, 94, 111, 52, 24, 201, 139, 228, 62, 196, 22, 164, 25, 52, 14, \\r
+94, 249, 3, 19, 114, 39, 40, 167, 109, 193, 29, 240, 40, 251, 246, 56, 1, \\r
+249, 188, 102, 248, 110, 209, 181, 56, 67, 182, 104, 172, 9, 253, 9, 104, \\r
+224, 34, 175, 64, 38, 12, 82, 167, 173, 195, 130, 220, 60, 101, 242, 232, \\r
+192, 6, 100, 177, 32, 149, 34, 221, 173, 255, 27, 139, 182, 74, 187, 152, 26, \\r
+236, 123, 166, 44, 225, 199, 220, 25, 6, 10, 159, 82, 43, 202, 47, 64, 248, \\r
+234, 135, 137, 253, 93, 0, 110, 0, 13, 48, 67, 90, 143, 138, 164, 80, 143, \\r
+12, 119, 242, 101, 228, 45, 229, 168, 127, 120, 227, 149, 218, 7, 41, 81, \\r
+148, 228, 1, 191, 110, 152, 156, 144, 168, 10, 61, 60, 80, 235, 18, 196, 95, \\r
+71, 123, 44, 223, 120, 19, 116, 224, 123, 113, 71, 72, 27, 227, 107, 11, 75, \\r
+89, 147, 223, 225, 27, 208, 237, 71, 77, 194, 222, 92, 218, 149, 200, 16, \\r
+119, 127, 192, 167, 211, 211, 216, 52, 161, 21, 224, 4, 30, 130, 202, 217, \\r
+17, 177, 201, 76, 75, 9, 72, 252, 221, 110, 202, 118, 174, 142, 59, 251, 101, \\r
+131, 29, 219, 174, 79, 198, 30, 197, 204, 124, 40, 249, 210, 134, 21, 250, \\r
+65, 69, 46, 149, 191, 13, 139, 141, 17, 156, 184, 251, 38, 214, 145, 13, 205, \\r
+7, 80, 250, 41, 63, 12, 43, 254, 151, 85, 90, 109, 150, 5, 9, 69, 239, 15, \\r
+116, 147, 108, 130, 55, 235, 80, 245, 196, 11, 56, 45, 138, 147, 109, 29, \\r
+229, 247, 231, 215, 92, 180, 80, 117, 13, 243, 206, 49, 241, 142, 143, 219, \\r
+20, 246, 3, 171, 174, 188, 235, 200, 103, 204, 255, 79, 198, 41, 149, 188, \\r
+101, 82, 242, 53, 196, 244, 155, 69, 28, 219, 97, 254, 63, 185, 216, 241, \\r
+172, 7, 164, 184, 136, 149, 184, 86, 50, 90, 62, 34, 11, 150, 238, 111, 30, \\r
+50, 73, 79, 204, 219, 81, 77, 209, 79, 9, 90, 172, 231, 134, 88, 165, 167, \\r
+94, 54, 17, 58, 137, 130, 220, 98, 14, 171, 43, 37, 249, 119, 134, 29, 121, \\r
+81, 189, 152, 183, 187, 246, 217, 102, 1, 18, 235, 10, 205, 26, 170, 204, \\r
+101, 66, 143, 42, 72, 198, 186, 216, 45, 41, 232, 87, 119, 122, 180, 73, 101, \\r
+67, 87, 254, 198, 199, 98, 103, 88, 214, 154, 200, 232, 48, 96, 115, 137, 35, \\r
+8, 122, 172, 93, 109, 206, 12, 27, 23, 193, 66, 83, 21, 124, 116, 43, 37, 86, \\r
+97, 236, 11, 66, 4, 202, 9, 171, 27, 208, 159, 106, 127, 65, 248, 234, 130, \\r
+129, 65, 206, 206, 13, 43, 51, 163, 100, 250, 166, 57, 127, 246, 178, 85, \\r
+107, 172, 51, 147, 231, 190, 52, 207, 239, 162, 238, 237, 203, 176, 137, 218, \\r
+59, 190, 62, 51, 81, 165, 60, 204, 160, 253, 57, 203, 252, 31, 141, 104, 212, \\r
+140, 170, 142, 58, 151, 38, 107, 172, 138, 75, 116, 63, 183, 191, 12, 230, \\r
+205, 66, 250, 238, 249, 12, 44, 121, 126, 216, 64, 124, 56, 73, 113, 117, \\r
+183, 54, 7, 26, 44, 30, 60, 122, 56, 74, 52, 32, 149, 96, 247, 195, 177, 110, \\r
+177, 112, 116, 59, 13, 109, 251, 229, 226, 102, 104, 8, 251, 60, 140, 221, 2, \\r
+46, 110, 241, 175, 121, 103, 226, 76, 165, 121, 198, 181, 175, 21, 254, 226, \\r
+43, 202, 230, 62, 44, 88, 114, 206, 113, 140, 190, 32, 102, 7, 111, 209, 130, \\r
+49, 243, 149, 145, 18, 243, 79, 108, 166, 142, 178, 120, 222, 121, 197, 34, \\r
+243, 7, 231, 158, 146, 115, 60, 161, 218, 33, 35, 69, 229, 116, 132, 194, \\r
+104, 53, 154, 200, 46, 29, 209, 166, 223, 6, 127, 140, 40, 246, 195, 248, 98, \\r
+249, 87, 39, 209, 211, 11, 163, 135, 131, 201, 20, 180, 197, 237, 84, 225, \\r
+185, 41, 111, 227, 13, 204, 246, 13, 36, 69, 110, 228, 219, 249, 193, 212, \\r
+180, 181, 102, 28, 110, 175, 64, 230, 73, 126, 65, 24, 135, 223, 187, 96, 5, \\r
+244, 93, 161, 214, 103, 42, 200, 119, 249, 138, 249, 38, 239, 116, 88, 106, \\r
+200, 53, 16, 197, 30, 11, 131, 106, 68, 141, 92, 251, 254, 146, 3, 67, 5, \\r
+175, 192, 243, 4, 40, 157, 9, 53, 227, 95, 69, 171, 205, 140, 253, 44, 157, \\r
+175, 100, 67, 128, 130, 229, 142, 61, 45, 219, 223, 171, 8, 22, 236, 126, \\r
+167, 203, 216, 92, 44, 136, 152, 194, 173, 29, 176, 177, 206, 147, 226, 244, \\r
+219, 165, 182, 73, 135, 29, 162, 121, 53, 190, 85, 205, 99, 44, 16, 31, 236, \\r
+59, 138, 63, 56, 128, 28, 45, 146, 95, 225, 181, 188, 131, 182, 27, 237, 79, \\r
+58, 196, 254, 255, 203, 129, 103, 194, 89, 218, 10, 156, 34, 235, 147, 226, \\r
+127, 172, 74, 130, 129, 63, 235, 139, 15, 156, 65, 130, 155, 124, 81, 167, \\r
+171, 173, 146, 25, 240, 253, 147, 246, 173, 36, 203, 222, 181, 171, 129, 84, \\r
+174, 36, 100, 210, 208, 254, 51, 42, 117, 7, 143, 58, 18, 159, 190, 27, 27, \\r
+31, 199, 241, 124, 145, 93, 50, 217, 214, 234, 6, 91, 105, 107, 172, 9, 22, \\r
+163, 202, 142, 131, 52, 183, 65, 39, 179, 1, 230, 6, 140, 4, 61, 190, 44, \\r
+102, 181, 134, 17, 19, 109, 208, 190, 141, 199, 85, 2, 193, 118, 95, 32, 94, \\r
+228, 62, 119, 225, 63, 232, 108, 74, 160, 254, 206, 63, 103, 148, 63, 239, \\r
+239, 134, 217, 74, 8, 139, 200, 28, 173, 32, 167, 140, 6, 103, 175, 115, 205, \\r
+22, 132, 50, 208, 243, 220, 219, 183, 36, 106, 99, 28, 161, 222, 188, 146, \\r
+241, 92, 231, 224, 213, 164, 74, 0, 214, 34, 156, 82, 143, 232, 34, 162, 240, \\r
+152, 56, 13, 246, 67, 121, 130, 175, 213, 193, 130, 82, 219, 143, 126, 209, \\r
+157, 166, 142, 158, 169, 142, 209, 246, 176, 234, 131, 233, 247, 59, 74, 9, \\r
+30, 50, 210, 93, 215, 171, 77, 111, 89, 252, 67, 65, 0, 60, 148, 82, 89, 108, \\r
+110, 232, 102, 245, 156, 87, 108, 226, 59, 162, 33, 238, 246, 79, 243, 98, \\r
+219, 143, 100, 211, 128, 105, 151, 71, 139, 182, 56, 69, 190, 13, 97, 207, \\r
+120, 207, 65, 127, 244, 90, 139, 114, 230, 125, 76, 24, 238, 58, 18, 148, 52, \\r
+48, 57, 0, 149, 96, 250, 127, 166, 199, 177, 126, 13, 36, 255, 67, 120, 136, \\r
+108, 7, 115, 146, 200, 11, 80, 151, 244, 27, 219, 69, 98, 145, 135, 93, 245, \\r
+97, 56, 243, 138, 79, 219, 61, 67, 47, 186, 169, 180, 169, 38, 126, 140, 22, \\r
+252, 28, 167, 157, 124, 133, 167, 53, 134, 79, 200, 154, 13, 36, 4, 88, 66, \\r
+48, 117, 141, 188, 114, 52, 249, 70, 94, 244, 92, 103, 4, 136, 40, 169, 236, \\r
+236, 124, 70, 110, 28, 217, 192, 97, 88, 94, 55, 176, 153, 151, 160, 155, \\r
+119, 74, 208, 181, 29, 166, 74, 197, 58, 148, 180, 80, 6, 192, 168, 2, 95, \\r
+14, 33, 3, 177, 150, 2, 151, 128, 26, 11, 61, 33, 127, 3, 41, 50, 102, 191, \\r
+59, 87, 83, 231, 213, 229, 69, 163, 26, 27, 248, 207, 179, 149, 89, 78, 100, \\r
+67, 244, 37, 204, 97, 53, 223, 54, 85, 207, 154, 17, 246, 212, 79, 155, 175, \\r
+242, 3, 25, 155, 171, 196, 25, 161, 234, 27, 37, 173, 170, 0, 91, 157, 111, \\r
+136, 111, 48, 133, 198, 23, 87, 17, 171, 2, 221, 146, 77, 130, 2, 53, 203, \\r
+68, 164, 65, 213, 18, 231, 77}\r
+\r
+e riformattato per importazione\r
+\r
+〈x8,x1〉;〈x5,xE〉;〈x7,x6〉;〈xD,x1〉;〈x7,x5〉;〈x1,x0〉;〈x8,x5〉;〈x9,x0〉;〈x1,xF〉;〈x1,x2〉;〈xE,x2〉;〈x3,xC〉;〈x1,xD〉;〈x0,x6〉;〈x3,xC〉;〈xD,x1〉;\r
+〈x8,x3〉;〈xE,xB〉;〈x7,x2〉;〈x1,xF〉;〈x9,x0〉;〈xF,x0〉;〈x4,x7〉;〈xA,x3〉;〈xD,x0〉;〈x4,x6〉;〈xC,x5〉;〈x3,x2〉;〈xC,x9〉;〈x0,xB〉;〈x1,xB〉;〈x7,x3〉;\r
+〈x1,x1〉;〈xA,xF〉;〈xD,x2〉;〈xB,xF〉;〈x0,xB〉;〈xB,xC〉;〈x0,x8〉;〈xE,x5〉;〈xF,xE〉;〈xA,xC〉;〈xF,xC〉;〈x9,x4〉;〈x3,x4〉;〈x8,x8〉;〈x2,x8〉;〈xD,xB〉;\r
+〈xD,x3〉;〈x3,xD〉;〈x3,x8〉;〈x6,xA〉;〈x4,xD〉;〈x7,x4〉;〈x4,x9〉;〈x7,xE〉;〈x6,xF〉;〈x4,xD〉;〈x3,xD〉;〈x3,xA〉;〈xC,x2〉;〈xD,x3〉;〈x3,x6〉;〈x7,x0〉;\r
+〈x8,x0〉;〈xC,x9〉;〈x3,xB〉;〈x5,x2〉;〈x8,xF〉;〈x1,xE〉;〈x8,x4〉;〈x5,x2〉;〈x2,xD〉;〈xA,xF〉;〈x1,xB〉;〈x0,x1〉;〈x3,x5〉;〈x7,x2〉;〈x1,x0〉;〈x0,x7〉;\r
+〈x5,xD〉;〈xA,xB〉;〈xE,x0〉;〈x7,x6〉;〈xB,xD〉;〈x3,xC〉;〈x0,xD〉;〈xC,xE〉;〈xB,x7〉;〈x9,xD〉;〈xA,x9〉;〈xF,x3〉;〈xC,x1〉;〈x9,x3〉;〈x8,xC〉;〈x4,xE〉;\r
+〈xF,x8〉;〈xC,x3〉;〈x1,x3〉;〈x3,xA〉;〈xA,x2〉;〈xB,xD〉;〈x1,x0〉;〈x9,xB〉;〈x5,xC〉;〈xC,x8〉;〈xE,x5〉;〈xC,x8〉;〈xA,x6〉;〈xF,xB〉;〈x1,x3〉;〈xC,x8〉;\r
+〈x6,x4〉;〈x0,x9〉;〈xD,x6〉;〈xD,x7〉;〈x3,x9〉;〈xF,x9〉;〈xE,x4〉;〈x0,x3〉;〈x4,x9〉;〈xC,xE〉;〈x7,x1〉;〈x5,x7〉;〈x3,x4〉;〈x2,xE〉;〈xA,x3〉;〈x5,x3〉;\r
+〈xF,x9〉;〈x2,xA〉;〈xF,x0〉;〈xE,x1〉;〈x6,x6〉;〈x8,x9〉;〈x2,xE〉;〈x8,x6〉;〈xC,x7〉;〈x2,xD〉;〈x1,x2〉;〈xD,x5〉;〈x4,xA〉;〈x7,x9〉;〈x6,xD〉;〈x1,xB〉;\r
+〈x4,x1〉;〈x9,x4〉;〈x6,xC〉;〈x6,xD〉;〈xF,x4〉;〈x5,xD〉;〈x8,xB〉;〈xB,xE〉;〈xC,x8〉;〈xE,x5〉;〈xA,x4〉;〈xB,xA〉;〈xB,x1〉;〈x2,x1〉;〈x5,x0〉;〈x6,xB〉;\r
+〈x9,x5〉;〈x0,x6〉;〈x4,x5〉;〈xE,x0〉;〈x0,xF〉;〈x5,xA〉;〈xC,xE〉;〈xC,x8〉;〈x8,x2〉;〈x4,xF〉;〈xC,x2〉;〈xF,x5〉;〈xF,x8〉;〈x1,x3〉;〈xD,xA〉;〈x2,xB〉;\r
+〈x7,x9〉;〈xB,xF〉;〈xA,x4〉;〈x5,xA〉;〈xD,x2〉;〈x7,x6〉;〈x8,xD〉;〈x1,xF〉;〈x1,x5〉;〈x5,xA〉;〈xD,xE〉;〈x2,x3〉;〈x9,xD〉;〈x7,xE〉;〈x6,x7〉;〈x6,x3〉;\r
+〈x1,x1〉;〈xF,x5〉;〈x2,x4〉;〈xD,x6〉;〈x6,xA〉;〈x0,xC〉;〈x7,xF〉;〈x1,x2〉;〈xD,x6〉;〈xE,xE〉;〈xB,xA〉;〈x4,x4〉;〈x3,x9〉;〈x9,x2〉;〈x6,x6〉;〈xE,xA〉;\r
+〈x2,x3〉;〈x4,x4〉;〈xC,xF〉;〈x7,x4〉;〈x6,xB〉;〈x8,x3〉;〈x7,x7〉;〈x0,x4〉;〈x6,x4〉;〈xB,x3〉;〈x3,xC〉;〈x2,x6〉;〈xF,xC〉;〈xD,x5〉;〈x4,x1〉;〈xB,x7〉;\r
+〈x2,x4〉;〈xE,x0〉;〈xB,x4〉;〈x1,x3〉;〈x3,x2〉;〈x4,x5〉;〈x1,x8〉;〈xD,xB〉;〈x0,x0〉;〈xB,x6〉;〈x5,xF〉;〈x3,xA〉;〈xB,x7〉;〈x4,xF〉;〈xB,x4〉;〈xD,xF〉;\r
+〈xF,x4〉;〈x1,xA〉;〈x1,x0〉;〈xE,x9〉;〈xC,x5〉;〈x2,xC〉;〈xB,xC〉;〈x5,xB〉;〈x0,x5〉;〈x4,xA〉;〈x5,x2〉;〈xC,x0〉;〈x1,xF〉;〈x5,x9〉;〈xF,x2〉;〈xE,x5〉;\r
+\r
+〈x9,x7〉;〈x5,x1〉;〈xF,x7〉;〈x3,xC〉;〈x7,x7〉;〈x8,x1〉;〈x7,xC〉;〈xB,x0〉;〈x1,xD〉;〈x1,xA〉;〈x7,xB〉;〈x4,x7〉;〈x7,x9〉;〈xC,x2〉;〈x3,xF〉;〈xD,x3〉;\r
+〈xB,x6〉;〈xD,x7〉;〈xC,x9〉;〈x9,x2〉;〈x0,xD〉;〈x7,xF〉;〈x0,xB〉;〈x7,x5〉;〈x0,xE〉;〈x5,x1〉;〈x9,xF〉;〈x4,xC〉;〈xE,x5〉;〈xD,x0〉;〈xC,x1〉;〈x3,x9〉;\r
+〈x7,xF〉;〈xE,xA〉;〈x4,x1〉;〈x0,x5〉;〈x0,x8〉;〈x7,xE〉;〈x7,x6〉;〈x2,xF〉;〈x6,x2〉;〈x7,x1〉;〈xF,x8〉;〈x3,x8〉;〈xB,x6〉;〈x7,x0〉;〈xD,xD〉;〈xD,x0〉;\r
+〈xA,xD〉;〈x7,x8〉;〈x3,xD〉;〈x7,xB〉;〈xC,xC〉;〈x6,xD〉;〈x3,xF〉;〈xA,xC〉;〈xC,x7〉;〈x6,xE〉;〈xF,xF〉;〈x4,x5〉;〈xC,x1〉;〈x7,x4〉;〈xB,xF〉;〈x2,x9〉;\r
+〈x4,x9〉;〈xC,x5〉;〈x7,xC〉;〈x2,xD〉;〈xF,x2〉;〈xC,xC〉;〈xF,xA〉;〈x5,x8〉;〈xA,xC〉;〈x5,x8〉;〈x5,x1〉;〈x0,xE〉;〈x4,x8〉;〈x7,x0〉;〈x4,x1〉;〈xE,xD〉;\r
+〈x9,x5〉;〈xB,x4〉;〈x4,x7〉;〈xF,x1〉;〈x8,x1〉;〈xE,x4〉;〈x4,x0〉;〈x5,x4〉;〈x7,xB〉;〈xA,x1〉;〈xD,x2〉;〈x0,x2〉;〈xC,x5〉;〈x7,xD〉;〈x7,xA〉;〈xF,x1〉;\r
+〈x0,x0〉;〈x8,xD〉;〈x6,x6〉;〈xB,x0〉;〈xA,x9〉;〈x7,xA〉;〈x2,xB〉;〈x2,xC〉;〈x6,xF〉;〈x5,xF〉;〈x6,xF〉;〈x4,xE〉;〈x2,xB〉;〈x1,x1〉;〈xF,xF〉;〈x4,x7〉;\r
+〈x7,xF〉;〈xC,xF〉;〈xD,x3〉;〈x8,x5〉;〈xB,x8〉;〈x4,xD〉;〈x9,x9〉;〈x8,x3〉;〈x7,x6〉;〈xE,x8〉;〈x3,x0〉;〈xF,x8〉;〈xD,x0〉;〈x5,x9〉;〈xB,x8〉;〈x7,x1〉;\r
+〈x5,x7〉;〈xA,xF〉;〈xB,x5〉;〈xD,x5〉;〈xF,xC〉;〈x8,x2〉;〈x1,x6〉;〈x1,x0〉;〈xB,x3〉;〈xC,x1〉;〈x1,xA〉;〈x0,x1〉;〈x1,x1〉;〈xF,xF〉;〈x5,x9〉;〈x2,x4〉;\r
+〈xC,x5〉;〈x7,x2〉;〈xD,x8〉;〈xD,x0〉;〈xA,xD〉;〈xF,xE〉;〈xD,x7〉;〈x1,x1〉;〈xC,xE〉;〈xF,x9〉;〈xB,x8〉;〈x7,x7〉;〈x3,xA〉;〈x1,xF〉;〈x6,x1〉;〈x1,xB〉;\r
+〈xB,xB〉;〈x5,x5〉;〈x8,x0〉;〈x7,xC〉;〈x2,x5〉;〈x3,x4〉;〈x8,x9〉;〈xF,x2〉;〈xC,x9〉;〈xD,xF〉;〈x3,x5〉;〈xC,x5〉;〈x1,x2〉;〈xF,x0〉;〈x0,x5〉;〈xD,xE〉;\r
+〈x2,x6〉;〈x4,x9〉;〈xB,x7〉;〈x3,x9〉;〈x0,x5〉;〈xC,x2〉;〈xD,xB〉;〈xF,xC〉;〈x9,xF〉;〈xA,x9〉;〈x6,x6〉;〈xA,xD〉;〈x4,xA〉;〈x3,xF〉;〈xB,xF〉;〈x6,xD〉;\r
+〈x8,x7〉;〈x6,xA〉;〈xB,x1〉;〈x3,xE〉;〈xB,x6〉;〈x0,xE〉;〈x7,xA〉;〈x3,xB〉;〈x4,x5〉;〈xE,x9〉;〈xC,xE〉;〈x6,xA〉;〈x6,xA〉;〈x7,x0〉;〈x6,x0〉;〈x6,xA〉;\r
+〈x2,xC〉;〈xD,x2〉;〈xB,x8〉;〈x3,x6〉;〈x2,x1〉;〈x0,x0〉;〈x5,x4〉;〈x3,x1〉;〈x6,x0〉;〈x1,xB〉;〈x4,xC〉;〈xC,xA〉;〈xB,xE〉;〈x5,xF〉;〈x8,x1〉;〈xB,x7〉;\r
+〈x9,xB〉;〈x2,x6〉;〈x9,x4〉;〈x2,xB〉;〈x4,x1〉;〈x2,xB〉;〈x9,x8〉;〈x6,x3〉;〈x6,x6〉;〈x6,x5〉;〈x4,x6〉;〈x2,x3〉;〈xE,x5〉;〈x0,x7〉;〈x9,xE〉;〈x1,xC〉;\r
+〈x3,x8〉;〈x5,xC〉;〈x9,x7〉;〈x6,x3〉;〈x5,x3〉;〈x6,x6〉;〈x0,x8〉;〈x5,xD〉;〈x0,x8〉;〈xD,xB〉;〈x6,xE〉;〈x5,x6〉;〈x7,x0〉;〈x3,x2〉;〈x4,x5〉;〈x0,x2〉;\r
+\r
+〈x6,x3〉;〈x7,x2〉;〈x9,xC〉;〈xD,x9〉;〈x5,x0〉;〈x0,x6〉;〈x5,x9〉;〈x1,x7〉;〈x6,x8〉;〈xD,x2〉;〈xD,x7〉;〈x8,xE〉;〈x6,x9〉;〈x5,xF〉;〈x8,x1〉;〈x8,x4〉;\r
+〈x8,x7〉;〈xD,xC〉;〈x9,x8〉;〈xE,x5〉;〈xB,x5〉;〈xC,x3〉;〈x2,x5〉;〈x6,xC〉;〈x9,x2〉;〈xD,xD〉;〈x2,xA〉;〈xD,x1〉;〈x1,x4〉;〈x7,xE〉;〈x1,x7〉;〈xB,x2〉;\r
+〈x9,x8〉;〈x5,x5〉;〈xF,xC〉;〈x3,xD〉;〈x8,xD〉;〈xE,xF〉;〈x8,x1〉;〈xB,x8〉;〈xB,xB〉;〈x5,x1〉;〈x0,x0〉;〈xB,x4〉;〈x2,xE〉;〈x3,x0〉;〈x6,x0〉;〈x7,xE〉;\r
+〈x9,x0〉;〈xE,x3〉;〈xF,x4〉;〈x7,x2〉;〈x1,xC〉;〈xB,x3〉;〈x7,x8〉;〈x1,xB〉;〈x9,xF〉;〈x1,xB〉;〈x0,x3〉;〈xA,x3〉;〈x0,x5〉;〈xD,xE〉;〈x3,x8〉;〈xB,xA〉;\r
+〈x0,xE〉;〈xE,xD〉;〈xE,xC〉;〈x1,xF〉;〈x3,x8〉;〈xE,x3〉;〈xF,x7〉;〈xA,xA〉;〈xE,x9〉;〈x3,xD〉;〈xF,xF〉;〈xF,x3〉;〈x1,x4〉;〈x2,xC〉;〈x8,x8〉;〈x6,x1〉;\r
+〈x3,x0〉;〈xA,xB〉;〈x1,x8〉;〈xD,xC〉;〈xF,xE〉;〈x6,xA〉;〈x2,x9〉;〈xF,x1〉;〈xC,xB〉;〈x9,x0〉;〈x7,x8〉;〈x9,x9〉;〈x1,xF〉;〈x2,x8〉;〈xF,x9〉;〈xC,xB〉;\r
+〈x1,x4〉;〈x8,x4〉;〈xF,x3〉;〈xD,x6〉;〈x7,xE〉;〈xE,xC〉;〈x5,x6〉;〈xC,xE〉;〈xD,xA〉;〈x5,xE〉;〈x6,x1〉;〈xF,x1〉;〈x6,x6〉;〈x6,x9〉;〈x9,x3〉;〈x5,x9〉;\r
+〈x3,xC〉;〈x1,xD〉;〈x6,xB〉;〈xF,x4〉;〈x5,x9〉;〈x4,xD〉;〈x3,x8〉;〈xA,x9〉;〈x3,xB〉;〈x7,xF〉;〈xB,x2〉;〈xE,xC〉;〈xA,xE〉;〈xF,x6〉;〈xB,x2〉;〈x2,x2〉;\r
+〈x6,x4〉;〈x2,x7〉;〈x6,xC〉;〈x2,x0〉;〈xE,xE〉;〈x5,x1〉;〈x3,xE〉;〈x8,x8〉;〈xD,xD〉;〈xC,x1〉;〈xD,xC〉;〈xC,x1〉;〈x6,x6〉;〈x6,x1〉;〈x4,x2〉;〈x7,x7〉;\r
+〈x3,x6〉;〈x0,x8〉;〈x2,x9〉;〈x6,x0〉;〈xA,x9〉;〈xF,xC〉;〈x7,xC〉;〈xA,x7〉;〈xB,x4〉;〈xF,xC〉;〈x8,x7〉;〈x1,xD〉;〈x6,xC〉;〈xA,x2〉;〈x3,xF〉;〈x1,xD〉;\r
+〈x1,x7〉;〈x0,xF〉;〈x0,x2〉;〈x2,x6〉;〈xA,x2〉;〈x6,xA〉;〈x5,xC〉;〈xE,xD〉;〈x2,x7〉;〈xC,x5〉;〈x7,xB〉;〈xF,x5〉;〈x9,xC〉;〈x8,x5〉;〈x6,x3〉;〈x5,x6〉;\r
+〈xC,x3〉;〈x4,xB〉;〈x1,xB〉;〈xA,x0〉;〈x1,xB〉;〈x8,x9〉;〈x3,x5〉;〈xD,x6〉;〈xD,x9〉;〈xD,xD〉;〈x2,xE〉;〈x6,x2〉;〈x7,x5〉;〈xE,x7〉;〈x1,x8〉;〈x4,xD〉;\r
+〈xD,x7〉;〈x5,x8〉;〈xA,x7〉;〈x5,xF〉;〈x9,x4〉;〈x8,x7〉;〈xA,x8〉;〈xE,x7〉;〈x2,xB〉;〈xF,x2〉;〈xE,x7〉;〈xB,x9〉;〈x0,x6〉;〈xA,xF〉;〈xD,xA〉;〈xD,xC〉;\r
+〈xC,x6〉;〈x3,xF〉;〈x8,xD〉;〈x7,x9〉;〈x9,x5〉;〈xD,xA〉;〈x5,xB〉;〈x9,x2〉;〈xE,xE〉;〈x3,xC〉;〈xF,xE〉;〈x4,x9〉;〈x5,xA〉;〈x1,x0〉;〈x4,xD〉;〈x8,x9〉;\r
+〈x8,x3〉;〈x2,x6〉;〈xE,xC〉;〈x8,xD〉;〈xC,x9〉;〈x7,x7〉;〈xE,xE〉;〈xF,x1〉;〈x4,x0〉;〈x6,xD〉;〈x4,x9〉;〈x5,x7〉;〈x9,xB〉;〈xC,x4〉;〈x1,xF〉;〈x8,x0〉;\r
+〈x9,x5〉;〈xB,xC〉;〈xE,x8〉;〈xF,x9〉;〈xD,x7〉;〈x1,x4〉;〈x3,xE〉;〈xC,x3〉;〈x6,xF〉;〈x8,xF〉;〈x7,x2〉;〈xD,x5〉;〈xB,xE〉;〈x8,xA〉;〈xA,x3〉;〈xF,x7〉;\r
+\r
+〈x6,x0〉;〈x3,xA〉;〈x7,x4〉;〈xF,xB〉;〈xB,xD〉;〈x7,x4〉;〈x8,x3〉;〈xE,x3〉;〈x9,xD〉;〈xD,x9〉;〈xB,x8〉;〈x1,x3〉;〈x5,x0〉;〈x4,x0〉;〈x8,xA〉;〈x9,x6〉;\r
+〈x3,xA〉;〈xA,x6〉;〈xE,xC〉;〈x7,xC〉;〈x1,x5〉;〈x8,x7〉;〈x4,xD〉;〈x6,xA〉;〈xA,xA〉;〈xE,x0〉;〈xB,xA〉;〈xF,xF〉;〈x3,xB〉;〈xE,x2〉;〈x5,x1〉;〈x2,x2〉;\r
+〈x2,x2〉;〈x1,xF〉;〈xA,x1〉;〈x2,x1〉;〈xA,xF〉;〈x3,x7〉;〈x8,xA〉;〈xD,xF〉;〈xE,x3〉;〈x6,x9〉;〈xE,xE〉;〈xC,x4〉;〈xE,x7〉;〈x7,x1〉;〈x9,x6〉;〈x1,x1〉;\r
+〈xE,x4〉;〈x3,x9〉;〈xE,x5〉;〈xA,xF〉;〈xF,x5〉;〈x5,x7〉;〈xE,xB〉;〈x5,x5〉;〈x6,x5〉;〈x8,xB〉;〈x3,xE〉;〈x8,xD〉;〈x4,x6〉;〈x5,x3〉;〈xB,x2〉;〈x1,x9〉;\r
+〈x3,x4〉;〈xE,x9〉;〈x4,xA〉;〈x4,xB〉;〈x5,x2〉;〈x3,x0〉;〈x3,xF〉;〈xA,x7〉;〈x4,xF〉;〈x1,xA〉;〈xB,x8〉;〈x6,x4〉;〈x5,xB〉;〈xD,x9〉;〈x6,xD〉;〈x6,x1〉;\r
+〈xA,x5〉;〈xC,xF〉;〈x8,xC〉;〈xD,xD〉;〈xE,x6〉;〈xD,x5〉;〈x3,x6〉;〈x0,xC〉;〈x8,xD〉;〈xF,x7〉;〈x4,xE〉;〈x9,xC〉;〈xB,xF〉;〈x2,xB〉;〈x4,x4〉;〈xD,x1〉;\r
+〈xC,x0〉;〈x8,x0〉;〈x0,x8〉;〈xA,xD〉;〈xC,xE〉;〈xB,xD〉;〈x4,xC〉;〈x5,x3〉;〈x6,x5〉;〈xB,x6〉;〈x4,x8〉;〈xF,x6〉;〈x6,x4〉;〈x7,xC〉;〈x9,x8〉;〈x1,x0〉;\r
+〈x9,xD〉;〈xF,xD〉;〈x4,x9〉;〈xC,x4〉;〈xD,xD〉;〈x1,x4〉;〈xB,x6〉;〈x6,xF〉;〈x3,xB〉;〈x4,x6〉;〈xD,x7〉;〈x1,xA〉;〈x4,x4〉;〈xA,x4〉;〈x8,x1〉;〈x3,x1〉;\r
+〈xA,x2〉;〈x4,x0〉;〈x7,x0〉;〈x3,x9〉;〈x9,xA〉;〈x4,xC〉;〈x4,xF〉;〈x9,x3〉;〈x9,xD〉;〈xD,x4〉;〈x9,x7〉;〈x3,x9〉;〈xA,x8〉;〈xA,x8〉;〈xF,x9〉;〈xB,x3〉;\r
+〈xE,x7〉;〈xD,x8〉;〈x4,xD〉;〈x6,xD〉;〈x8,x8〉;〈x6,xB〉;〈x5,x4〉;〈x5,x5〉;〈x9,x2〉;〈x1,x9〉;〈xE,x4〉;〈xB,x1〉;〈xE,x3〉;〈x4,x6〉;〈xD,x5〉;〈x6,xC〉;\r
+〈xF,x9〉;〈x3,x3〉;〈xE,x5〉;〈xD,x2〉;〈x7,x7〉;〈x1,x8〉;〈x8,x6〉;〈x1,x4〉;〈x7,xE〉;〈x1,xE〉;〈xC,x8〉;〈xB,x4〉;〈xC,xE〉;〈xC,x7〉;〈x5,x7〉;〈x2,xD〉;\r
+〈x5,x2〉;〈xE,x7〉;〈x5,x7〉;〈x9,xB〉;〈x1,xA〉;〈x4,x9〉;〈x0,x9〉;〈x4,xB〉;〈xF,x6〉;〈xB,x5〉;〈x0,xA〉;〈x2,xC〉;〈xB,xA〉;〈x1,xF〉;〈x2,x6〉;〈x3,x8〉;\r
+〈x6,x0〉;〈x3,x4〉;〈xC,x4〉;〈x8,x9〉;〈xB,xF〉;〈x0,x4〉;〈x5,xF〉;〈xF,x4〉;〈x1,x1〉;〈xF,x3〉;〈x4,xA〉;〈xE,x2〉;〈xD,x8〉;〈x3,x6〉;〈xF,xA〉;〈x4,x3〉;\r
+〈x0,xC〉;〈x6,x0〉;〈x5,x1〉;〈x2,xD〉;〈x0,x3〉;〈xF,xC〉;〈x9,x2〉;〈x2,x1〉;〈xC,xA〉;〈x7,x5〉;〈x6,x6〉;〈xD,xC〉;〈x8,x4〉;〈xC,x4〉;〈x6,xF〉;〈xF,x0〉;\r
+〈x0,x4〉;〈x5,x2〉;〈xA,xB〉;〈x9,xA〉;〈xC,xE〉;〈xA,xA〉;〈x2,xF〉;〈x3,x6〉;〈xF,x3〉;〈xA,x7〉;〈x4,x2〉;〈x0,x7〉;〈x5,xE〉;〈xB,x6〉;〈x5,xB〉;〈x9,x8〉;\r
+〈x7,x9〉;〈x4,x3〉;〈x8,x8〉;〈x9,xA〉;〈x0,xA〉;〈x5,x5〉;〈xE,x3〉;〈x9,x8〉;〈x7,x5〉;〈xA,x2〉;〈xE,xA〉;〈xB,x9〉;〈x3,x2〉;〈x4,x1〉;〈x0,x7〉;〈x3,x8〉;\r
+\r
+〈xB,x1〉;〈x0,x3〉;〈x4,xE〉;〈x6,xF〉;〈x9,x7〉;〈x3,xC〉;〈x5,x4〉;〈xB,xB〉;〈xB,xC〉;〈x0,xB〉;〈x0,xC〉;〈xA,xB〉;〈x3,xB〉;〈x2,x8〉;〈xA,xA〉;〈x0,x4〉;\r
+〈x5,x0〉;〈xA,x3〉;〈x6,xB〉;〈xF,xA〉;〈xF,x0〉;〈xE,x6〉;〈x1,x8〉;〈xD,x8〉;〈x7,xA〉;〈x5,xF〉;〈x0,x3〉;〈x8,x2〉;〈x9,x6〉;〈x9,xA〉;〈xB,xF〉;〈x1,x5〉;\r
+〈x8,x5〉;〈xE,xF〉;〈xB,x3〉;〈x7,xB〉;〈xE,xE〉;〈xE,xF〉;〈x1,x0〉;〈x6,x9〉;〈xC,x2〉;〈xF,x9〉;〈x2,xD〉;〈x8,x1〉;〈xF,x3〉;〈x0,xE〉;〈xC,x3〉;〈x7,xF〉;\r
+〈xD,xD〉;〈x2,xE〉;〈x2,xA〉;〈xB,xD〉;〈xE,xA〉;〈x9,x2〉;〈xF,xF〉;〈xF,xA〉;〈x7,xB〉;〈xD,x3〉;〈x3,x0〉;〈x7,x5〉;〈x6,x7〉;〈xA,x8〉;〈x0,xF〉;〈x2,x1〉;\r
+〈xD,xC〉;〈xE,x5〉;〈xE,x2〉;〈x8,x9〉;〈x2,x9〉;〈xC,x5〉;〈xA,x3〉;〈xA,x2〉;〈x4,x2〉;〈x3,xF〉;〈xA,x3〉;〈x5,x8〉;〈xE,x0〉;〈x7,xC〉;〈x0,x3〉;〈xF,xF〉;\r
+〈x2,x8〉;〈x8,xB〉;〈x8,xB〉;〈x1,x2〉;〈xD,x8〉;〈xA,x8〉;〈x7,x6〉;〈xB,x9〉;〈xE,x2〉;〈xF,xE〉;〈x2,x1〉;〈x3,xF〉;〈xA,xC〉;〈x4,x6〉;〈xB,xC〉;〈xF,x8〉;\r
+〈xD,x3〉;〈xE,xB〉;〈xF,xC〉;〈x9,xF〉;〈xE,x7〉;〈x6,x1〉;〈xC,xB〉;〈xB,xF〉;〈x4,xE〉;〈xC,x4〉;〈x9,x7〉;〈x1,xE〉;〈x0,xD〉;〈x7,x9〉;〈x8,x3〉;〈xA,xB〉;\r
+〈x4,xC〉;〈x2,x6〉;〈x6,x3〉;〈x6,xF〉;〈xE,xE〉;〈x5,x9〉;〈x8,x1〉;〈x0,x2〉;〈x2,xC〉;〈xE,xD〉;〈x6,xF〉;〈x0,x4〉;〈x1,x0〉;〈xE,x0〉;〈xD,xA〉;〈xB,xE〉;\r
+〈xE,xE〉;〈x5,x7〉;〈xB,x0〉;〈x3,x1〉;〈x4,x1〉;〈xD,xC〉;〈x3,xC〉;〈xC,xC〉;〈x5,x8〉;〈x2,x8〉;〈x2,xC〉;〈x1,xB〉;〈x8,x6〉;〈xD,x6〉;〈xF,x9〉;〈xD,x5〉;\r
+〈x4,xA〉;〈xE,xA〉;〈x0,xB〉;〈x2,x0〉;〈x2,xC〉;〈x4,x2〉;〈xC,xE〉;〈x4,x5〉;〈x2,xB〉;〈x0,x1〉;〈xA,xA〉;〈xB,x1〉;〈x6,xE〉;〈xB,x7〉;〈xB,x7〉;〈x2,x8〉;\r
+〈x9,x5〉;〈x1,x9〉;〈xA,x7〉;〈x5,xC〉;〈x4,xE〉;〈x7,xB〉;〈x3,xE〉;〈xD,x3〉;〈x9,x0〉;〈x8,x6〉;〈x7,x1〉;〈x1,x4〉;〈xD,x2〉;〈xD,x4〉;〈xB,x4〉;〈xF,x2〉;\r
+〈x3,x1〉;〈x2,x8〉;〈x4,x5〉;〈xF,xD〉;〈x7,x8〉;〈x5,xD〉;〈xF,xA〉;〈xF,x3〉;〈x9,x5〉;〈x4,xD〉;〈x3,x1〉;〈xB,x8〉;〈xC,xC〉;〈x2,x1〉;〈x1,x9〉;〈x4,x2〉;\r
+〈x2,xA〉;〈xF,x2〉;〈xB,xA〉;〈x0,x4〉;〈x9,xF〉;〈x4,x3〉;〈x4,x5〉;〈x1,xC〉;〈x7,x4〉;〈xB,xB〉;〈x7,x0〉;〈x5,xE〉;〈x0,x1〉;〈xA,xC〉;〈x6,xD〉;〈xD,x7〉;\r
+〈x9,xC〉;〈x9,xD〉;〈x1,xA〉;〈x9,x8〉;〈xB,x1〉;〈xF,xC〉;〈x6,x1〉;〈xA,x3〉;〈x4,x1〉;〈x4,x1〉;〈xA,xF〉;〈x1,xD〉;〈xE,x1〉;〈x3,x2〉;〈x1,x9〉;〈x6,x0〉;\r
+〈x2,x9〉;〈x9,x7〉;〈x8,x5〉;〈x5,x3〉;〈x5,x3〉;〈x9,x1〉;〈xB,x3〉;〈x9,x4〉;〈xD,x5〉;〈x9,xD〉;〈x4,xC〉;〈x3,x6〉;〈x0,xE〉;〈x8,x4〉;〈xA,x1〉;〈x4,x6〉;\r
+〈x6,xA〉;〈x1,xF〉;〈xF,x3〉;〈x6,xB〉;〈xB,xE〉;〈x4,xA〉;〈x1,x9〉;〈x7,x5〉;〈xF,xC〉;〈xC,x6〉;〈xE,xA〉;〈x7,xE〉;〈xD,x1〉;〈x3,x3〉;〈x6,x7〉;〈xB,x7〉;\r
+\r
+〈xE,xE〉;〈x5,x9〉;〈xE,x2〉;〈xD,xD〉;〈x2,x2〉;〈x8,xC〉;〈x9,xB〉;〈x3,xE〉;〈x9,x8〉;〈xF,xC〉;〈x1,x3〉;〈xE,x2〉;〈x0,xC〉;〈x4,xE〉;〈x3,x1〉;〈x8,x7〉;\r
+〈x6,x7〉;〈x6,xA〉;〈x4,xC〉;〈x4,xC〉;〈x7,x2〉;〈x0,x0〉;〈x0,x5〉;〈x1,xF〉;〈xF,x6〉;〈x3,x0〉;〈xE,xE〉;〈xD,xE〉;〈xB,x1〉;〈x4,xC〉;〈xF,x7〉;〈xE,xC〉;\r
+〈x2,xC〉;〈x4,x0〉;〈x6,xB〉;〈x6,x8〉;〈x9,x0〉;〈x8,x8〉;〈x6,xF〉;〈xB,x3〉;〈x4,x7〉;〈x6,x2〉;〈x9,x2〉;〈x9,xB〉;〈x2,xB〉;〈x3,x2〉;〈x4,x0〉;〈xA,x7〉;\r
+〈x8,x9〉;〈x4,x0〉;〈x2,x3〉;〈x5,xC〉;〈xF,x9〉;〈x2,x9〉;〈x6,x2〉;〈xA,xE〉;〈x5,xB〉;〈xC,x9〉;〈x2,xC〉;〈x9,x2〉;〈x6,xF〉;〈xF,x5〉;〈xA,x0〉;〈x0,xE〉;\r
+〈xD,xE〉;〈xF,x9〉;〈x0,x9〉;〈x1,x0〉;〈x3,x9〉;〈x4,x6〉;〈xC,x5〉;〈xE,x2〉;〈x8,x3〉;〈xD,x5〉;〈x8,xE〉;〈x4,x6〉;〈x4,xC〉;〈xA,xC〉;〈x7,xF〉;〈x4,xF〉;\r
+〈xC,x1〉;〈x4,xF〉;〈x1,xA〉;〈x6,x1〉;〈x9,x6〉;〈x0,xB〉;〈x0,x0〉;〈x6,xF〉;〈x2,x6〉;〈x8,xC〉;〈xE,xE〉;〈x9,x3〉;〈x1,xB〉;〈x9,xE〉;〈xA,x5〉;〈x9,x6〉;\r
+〈x2,xA〉;〈xE,xB〉;〈x4,x6〉;〈x5,xF〉;〈x3,xC〉;〈xD,x6〉;〈x2,xD〉;〈x9,x4〉;〈x6,xB〉;〈xF,x4〉;〈xD,xA〉;〈x6,x9〉;〈x5,x9〉;〈xA,xC〉;〈xB,xD〉;〈x9,xE〉;\r
+〈x4,x8〉;〈x0,x2〉;〈xD,xC〉;〈x5,xC〉;〈x6,x0〉;〈x2,xA〉;〈x6,xE〉;〈xC,xA〉;〈x6,xE〉;〈x1,xF〉;〈xD,x4〉;〈x3,xA〉;〈xB,x0〉;〈x9,xE〉;〈x8,xF〉;〈xA,xB〉;\r
+〈xB,x2〉;〈x0,x2〉;〈x4,x7〉;〈x7,xD〉;〈xA,xB〉;〈xD,xB〉;〈xB,x5〉;〈x6,xD〉;〈xE,x2〉;〈x8,x9〉;〈x4,xD〉;〈x0,x4〉;〈xB,xE〉;〈xF,xA〉;〈x2,x2〉;〈x1,x4〉;\r
+〈x7,x1〉;〈x1,x2〉;〈x1,xB〉;〈x0,xD〉;〈xB,xA〉;〈x5,xA〉;〈x6,xC〉;〈x1,xE〉;〈x3,xA〉;〈x0,xF〉;〈x6,xE〉;〈x4,x4〉;〈xC,x8〉;〈xB,x5〉;〈x8,xC〉;〈x0,x3〉;\r
+〈x0,x6〉;〈x6,x4〉;〈x8,x5〉;〈x2,x8〉;〈x6,x4〉;〈x2,x2〉;〈x8,x1〉;〈x7,x6〉;〈xF,xE〉;〈xF,xA〉;〈x6,x2〉;〈x9,x1〉;〈xB,xE〉;〈xB,xC〉;〈x6,x1〉;〈x4,xB〉;\r
+〈x7,xE〉;〈x5,x0〉;〈xB,xC〉;〈xE,xE〉;〈x6,x3〉;〈xC,xF〉;〈x1,xD〉;〈xF,xD〉;〈x6,x2〉;〈x5,xC〉;〈x8,x5〉;〈x9,xE〉;〈xA,x5〉;〈x2,x6〉;〈xE,x7〉;〈x4,x6〉;\r
+〈x3,xB〉;〈xE,xA〉;〈xB,xE〉;〈x0,x4〉;〈x8,x8〉;〈xF,x2〉;〈x9,x2〉;〈x0,xB〉;〈xD,x9〉;〈xE,x9〉;〈x2,x9〉;〈x3,x8〉;〈x8,x8〉;〈x8,xA〉;〈x6,x9〉;〈x1,x7〉;\r
+〈x4,xB〉;〈xB,xF〉;〈x0,xC〉;〈xF,x2〉;〈xF,xD〉;〈x7,x3〉;〈x5,x9〉;〈xB,xE〉;〈x5,x4〉;〈x1,xC〉;〈xD,x3〉;〈x3,x1〉;〈x6,x2〉;〈x1,xB〉;〈xB,x7〉;〈x3,x2〉;\r
+〈xA,x4〉;〈xF,x1〉;〈x7,x0〉;〈x9,xA〉;〈x4,x6〉;〈xA,x1〉;〈x1,xC〉;〈x0,x4〉;〈x6,xC〉;〈xF,x2〉;〈xE,x6〉;〈xC,x1〉;〈xA,x4〉;〈xF,x2〉;〈x2,xA〉;〈x4,xB〉;\r
+〈x3,x5〉;〈x9,xB〉;〈x9,x9〉;〈xF,xF〉;〈x0,x1〉;〈x1,x3〉;〈xF,x9〉;〈x5,xC〉;〈x3,xC〉;〈x5,x1〉;〈x8,xA〉;〈xA,x5〉;〈x5,xF〉;〈x9,xE〉;〈x5,xE〉;〈xC,x6〉;\r
+\r
+〈x2,x1〉;〈x3,x7〉;〈xD,x2〉;〈xB,x9〉;〈x9,x8〉;〈xA,x1〉;〈x6,x0〉;〈xE,x9〉;〈x4,x5〉;〈xC,xA〉;〈xD,x7〉;〈xB,xD〉;〈xC,xF〉;〈x0,xF〉;〈x2,x4〉;〈xE,x5〉;\r
+〈x7,x9〉;〈x4,xB〉;〈x1,xC〉;〈x5,x7〉;〈x3,xA〉;〈x2,x4〉;〈x2,x2〉;〈x0,x8〉;〈x3,x3〉;〈xE,x2〉;〈xA,x2〉;〈x5,x8〉;〈x2,x5〉;〈x5,x4〉;〈x7,x1〉;〈x2,xB〉;\r
+〈xF,xF〉;〈xE,xD〉;〈x4,x8〉;〈xF,x6〉;〈x2,x3〉;〈x3,x1〉;〈xB,xA〉;〈x5,x1〉;〈x9,xF〉;〈xA,xA〉;〈xC,xC〉;〈x0,x3〉;〈x1,x5〉;〈xC,x7〉;〈x2,xD〉;〈xD,x3〉;\r
+〈xE,xB〉;〈x8,xF〉;〈x8,x4〉;〈x4,x0〉;〈x5,x3〉;〈xA,xD〉;〈x6,x7〉;〈xE,xC〉;〈xA,xF〉;〈xD,xC〉;〈x1,xC〉;〈x7,x4〉;〈x6,xB〉;〈xA,xD〉;〈xC,xD〉;〈xA,x7〉;\r
+〈x1,x1〉;〈x1,x0〉;〈xC,xF〉;〈xB,xE〉;〈xA,x1〉;〈x0,x1〉;〈x3,xF〉;〈xC,x0〉;〈x8,x5〉;〈x2,x8〉;〈x6,xB〉;〈xC,x3〉;〈x6,xD〉;〈xD,x8〉;〈x7,x5〉;〈x5,xA〉;\r
+〈xF,x0〉;〈x2,x2〉;〈x4,xB〉;〈x9,xC〉;〈x3,x1〉;〈xE,x4〉;〈xE,x7〉;〈xC,x6〉;〈xF,xC〉;〈x3,x0〉;〈xD,x5〉;〈xF,x9〉;〈x1,xA〉;〈x4,x0〉;〈x1,xF〉;〈x6,xD〉;\r
+〈xD,x5〉;〈x7,x8〉;〈xB,x5〉;〈x7,x6〉;〈xC,x9〉;〈xE,x1〉;〈xD,xF〉;〈x1,x2〉;〈x6,x1〉;〈xD,xF〉;〈x9,xF〉;〈x5,x7〉;〈x7,xD〉;〈x0,xB〉;〈xA,xD〉;〈x5,xA〉;\r
+〈xA,x1〉;〈x8,x4〉;〈xE,x5〉;〈xF,x7〉;〈xB,xC〉;〈xD,x3〉;〈xA,x5〉;〈xB,x4〉;〈x8,x5〉;〈x6,x7〉;〈x3,x6〉;〈xF,xC〉;〈xB,x1〉;〈xB,x3〉;〈xC,xB〉;〈x1,xE〉;\r
+〈xE,xC〉;〈x6,xE〉;〈xE,x1〉;〈x1,xC〉;〈xA,x5〉;〈x5,x3〉;〈x9,x8〉;〈xF,x6〉;〈xD,xF〉;〈x4,x1〉;〈x1,x3〉;〈x2,xE〉;〈x7,xF〉;〈x0,xE〉;〈x3,x8〉;〈x3,xC〉;\r
+〈xD,x4〉;〈x8,xC〉;〈x2,xA〉;〈x2,x8〉;〈x4,xE〉;〈x7,xE〉;〈x0,xE〉;〈xF,x7〉;〈xC,xA〉;〈x3,xE〉;〈xE,x4〉;〈xB,x4〉;〈x0,x5〉;〈x5,x8〉;〈xD,xC〉;〈x7,x8〉;\r
+〈xD,x9〉;〈xF,x9〉;〈x7,x9〉;〈x8,x4〉;〈x0,x2〉;〈x3,xF〉;〈xC,xF〉;〈x3,x8〉;〈xD,x7〉;〈x2,x6〉;〈x1,xD〉;〈x1,x8〉;〈x4,xD〉;〈xE,xA〉;〈x7,xA〉;〈xD,x4〉;\r
+〈x2,x4〉;〈x0,xD〉;〈x4,xD〉;〈x9,x0〉;〈x1,x7〉;〈x1,xE〉;〈x6,xE〉;〈xB,x6〉;〈xC,xC〉;〈xC,x0〉;〈xB,x0〉;〈x5,xE〉;〈x9,x9〉;〈x6,xD〉;〈xC,xF〉;〈xE,xE〉;\r
+〈x2,x9〉;〈xC,xF〉;〈xA,x2〉;〈x0,xC〉;〈xA,xB〉;〈x7,x4〉;〈x2,x9〉;〈x4,xE〉;〈x8,x2〉;〈x9,x3〉;〈x6,x9〉;〈x7,xB〉;〈xE,xC〉;〈xC,x7〉;〈x8,x9〉;〈xC,xA〉;\r
+〈xD,xD〉;〈xA,xC〉;〈x6,x5〉;〈x8,x0〉;〈x1,x4〉;〈x0,x9〉;〈x5,x5〉;〈x0,xE〉;〈x8,x4〉;〈x5,xE〉;〈x6,xF〉;〈x3,x4〉;〈x1,x8〉;〈xC,x9〉;〈x8,xB〉;〈xE,x4〉;\r
+〈x3,xE〉;〈xC,x4〉;〈x1,x6〉;〈xA,x4〉;〈x1,x9〉;〈x3,x4〉;〈x0,xE〉;〈x5,xE〉;〈xF,x9〉;〈x0,x3〉;〈x1,x3〉;〈x7,x2〉;〈x2,x7〉;〈x2,x8〉;〈xA,x7〉;〈x6,xD〉;\r
+〈xC,x1〉;〈x1,xD〉;〈xF,x0〉;〈x2,x8〉;〈xF,xB〉;〈xF,x6〉;〈x3,x8〉;〈x0,x1〉;〈xF,x9〉;〈xB,xC〉;〈x6,x6〉;〈xF,x8〉;〈x6,xE〉;〈xD,x1〉;〈xB,x5〉;〈x3,x8〉;\r
+\r
+〈x4,x3〉;〈xB,x6〉;〈x6,x8〉;〈xA,xC〉;〈x0,x9〉;〈xF,xD〉;〈x0,x9〉;〈x6,x8〉;〈xE,x0〉;〈x2,x2〉;〈xA,xF〉;〈x4,x0〉;〈x2,x6〉;〈x0,xC〉;〈x5,x2〉;〈xA,x7〉;\r
+〈xA,xD〉;〈xC,x3〉;〈x8,x2〉;〈xD,xC〉;〈x3,xC〉;〈x6,x5〉;〈xF,x2〉;〈xE,x8〉;〈xC,x0〉;〈x0,x6〉;〈x6,x4〉;〈xB,x1〉;〈x2,x0〉;〈x9,x5〉;〈x2,x2〉;〈xD,xD〉;\r
+〈xA,xD〉;〈xF,xF〉;〈x1,xB〉;〈x8,xB〉;〈xB,x6〉;〈x4,xA〉;〈xB,xB〉;〈x9,x8〉;〈x1,xA〉;〈xE,xC〉;〈x7,xB〉;〈xA,x6〉;〈x2,xC〉;〈xE,x1〉;〈xC,x7〉;〈xD,xC〉;\r
+〈x1,x9〉;〈x0,x6〉;〈x0,xA〉;〈x9,xF〉;〈x5,x2〉;〈x2,xB〉;〈xC,xA〉;〈x2,xF〉;〈x4,x0〉;〈xF,x8〉;〈xE,xA〉;〈x8,x7〉;〈x8,x9〉;〈xF,xD〉;〈x5,xD〉;〈x0,x0〉;\r
+〈x6,xE〉;〈x0,x0〉;〈x0,xD〉;〈x3,x0〉;〈x4,x3〉;〈x5,xA〉;〈x8,xF〉;〈x8,xA〉;〈xA,x4〉;〈x5,x0〉;〈x8,xF〉;〈x0,xC〉;〈x7,x7〉;〈xF,x2〉;〈x6,x5〉;〈xE,x4〉;\r
+〈x2,xD〉;〈xE,x5〉;〈xA,x8〉;〈x7,xF〉;〈x7,x8〉;〈xE,x3〉;〈x9,x5〉;〈xD,xA〉;〈x0,x7〉;〈x2,x9〉;〈x5,x1〉;〈x9,x4〉;〈xE,x4〉;〈x0,x1〉;〈xB,xF〉;〈x6,xE〉;\r
+〈x9,x8〉;〈x9,xC〉;〈x9,x0〉;〈xA,x8〉;〈x0,xA〉;〈x3,xD〉;〈x3,xC〉;〈x5,x0〉;〈xE,xB〉;〈x1,x2〉;〈xC,x4〉;〈x5,xF〉;〈x4,x7〉;〈x7,xB〉;〈x2,xC〉;〈xD,xF〉;\r
+〈x7,x8〉;〈x1,x3〉;〈x7,x4〉;〈xE,x0〉;〈x7,xB〉;〈x7,x1〉;〈x4,x7〉;〈x4,x8〉;〈x1,xB〉;〈xE,x3〉;〈x6,xB〉;〈x0,xB〉;〈x4,xB〉;〈x5,x9〉;〈x9,x3〉;〈xD,xF〉;\r
+〈xE,x1〉;〈x1,xB〉;〈xD,x0〉;〈xE,xD〉;〈x4,x7〉;〈x4,xD〉;〈xC,x2〉;〈xD,xE〉;〈x5,xC〉;〈xD,xA〉;〈x9,x5〉;〈xC,x8〉;〈x1,x0〉;〈x7,x7〉;〈x7,xF〉;〈xC,x0〉;\r
+〈xA,x7〉;〈xD,x3〉;〈xD,x3〉;〈xD,x8〉;〈x3,x4〉;〈xA,x1〉;〈x1,x5〉;〈xE,x0〉;〈x0,x4〉;〈x1,xE〉;〈x8,x2〉;〈xC,xA〉;〈xD,x9〉;〈x1,x1〉;〈xB,x1〉;〈xC,x9〉;\r
+〈x4,xC〉;〈x4,xB〉;〈x0,x9〉;〈x4,x8〉;〈xF,xC〉;〈xD,xD〉;〈x6,xE〉;〈xC,xA〉;〈x7,x6〉;〈xA,xE〉;〈x8,xE〉;〈x3,xB〉;〈xF,xB〉;〈x6,x5〉;〈x8,x3〉;〈x1,xD〉;\r
+〈xD,xB〉;〈xA,xE〉;〈x4,xF〉;〈xC,x6〉;〈x1,xE〉;〈xC,x5〉;〈xC,xC〉;〈x7,xC〉;〈x2,x8〉;〈xF,x9〉;〈xD,x2〉;〈x8,x6〉;〈x1,x5〉;〈xF,xA〉;〈x4,x1〉;〈x4,x5〉;\r
+〈x2,xE〉;〈x9,x5〉;〈xB,xF〉;〈x0,xD〉;〈x8,xB〉;〈x8,xD〉;〈x1,x1〉;〈x9,xC〉;〈xB,x8〉;〈xF,xB〉;〈x2,x6〉;〈xD,x6〉;〈x9,x1〉;〈x0,xD〉;〈xC,xD〉;〈x0,x7〉;\r
+〈x5,x0〉;〈xF,xA〉;〈x2,x9〉;〈x3,xF〉;〈x0,xC〉;〈x2,xB〉;〈xF,xE〉;〈x9,x7〉;〈x5,x5〉;〈x5,xA〉;〈x6,xD〉;〈x9,x6〉;〈x0,x5〉;〈x0,x9〉;〈x4,x5〉;〈xE,xF〉;\r
+〈x0,xF〉;〈x7,x4〉;〈x9,x3〉;〈x6,xC〉;〈x8,x2〉;〈x3,x7〉;〈xE,xB〉;〈x5,x0〉;〈xF,x5〉;〈xC,x4〉;〈x0,xB〉;〈x3,x8〉;〈x2,xD〉;〈x8,xA〉;〈x9,x3〉;〈x6,xD〉;\r
+〈x1,xD〉;〈xE,x5〉;〈xF,x7〉;〈xE,x7〉;〈xD,x7〉;〈x5,xC〉;〈xB,x4〉;〈x5,x0〉;〈x7,x5〉;〈x0,xD〉;〈xF,x3〉;〈xC,xE〉;〈x3,x1〉;〈xF,x1〉;〈x8,xE〉;〈x8,xF〉;\r
+\r
+〈xD,xB〉;〈x1,x4〉;〈xF,x6〉;〈x0,x3〉;〈xA,xB〉;〈xA,xE〉;〈xB,xC〉;〈xE,xB〉;〈xC,x8〉;〈x6,x7〉;〈xC,xC〉;〈xF,xF〉;〈x4,xF〉;〈xC,x6〉;〈x2,x9〉;〈x9,x5〉;\r
+〈xB,xC〉;〈x6,x5〉;〈x5,x2〉;〈xF,x2〉;〈x3,x5〉;〈xC,x4〉;〈xF,x4〉;〈x9,xB〉;〈x4,x5〉;〈x1,xC〉;〈xD,xB〉;〈x6,x1〉;〈xF,xE〉;〈x3,xF〉;〈xB,x9〉;〈xD,x8〉;\r
+〈xF,x1〉;〈xA,xC〉;〈x0,x7〉;〈xA,x4〉;〈xB,x8〉;〈x8,x8〉;〈x9,x5〉;〈xB,x8〉;〈x5,x6〉;〈x3,x2〉;〈x5,xA〉;〈x3,xE〉;〈x2,x2〉;〈x0,xB〉;〈x9,x6〉;〈xE,xE〉;\r
+〈x6,xF〉;〈x1,xE〉;〈x3,x2〉;〈x4,x9〉;〈x4,xF〉;〈xC,xC〉;〈xD,xB〉;〈x5,x1〉;〈x4,xD〉;〈xD,x1〉;〈x4,xF〉;〈x0,x9〉;〈x5,xA〉;〈xA,xC〉;〈xE,x7〉;〈x8,x6〉;\r
+〈x5,x8〉;〈xA,x5〉;〈xA,x7〉;〈x5,xE〉;〈x3,x6〉;〈x1,x1〉;〈x3,xA〉;〈x8,x9〉;〈x8,x2〉;〈xD,xC〉;〈x6,x2〉;〈x0,xE〉;〈xA,xB〉;〈x2,xB〉;〈x2,x5〉;〈xF,x9〉;\r
+〈x7,x7〉;〈x8,x6〉;〈x1,xD〉;〈x7,x9〉;〈x5,x1〉;〈xB,xD〉;〈x9,x8〉;〈xB,x7〉;〈xB,xB〉;〈xF,x6〉;〈xD,x9〉;〈x6,x6〉;〈x0,x1〉;〈x1,x2〉;〈xE,xB〉;〈x0,xA〉;\r
+〈xC,xD〉;〈x1,xA〉;〈xA,xA〉;〈xC,xC〉;〈x6,x5〉;〈x4,x2〉;〈x8,xF〉;〈x2,xA〉;〈x4,x8〉;〈xC,x6〉;〈xB,xA〉;〈xD,x8〉;〈x2,xD〉;〈x2,x9〉;〈xE,x8〉;〈x5,x7〉;\r
+〈x7,x7〉;〈x7,xA〉;〈xB,x4〉;〈x4,x9〉;〈x6,x5〉;〈x4,x3〉;〈x5,x7〉;〈xF,xE〉;〈xC,x6〉;〈xC,x7〉;〈x6,x2〉;〈x6,x7〉;〈x5,x8〉;〈xD,x6〉;〈x9,xA〉;〈xC,x8〉;\r
+〈xE,x8〉;〈x3,x0〉;〈x6,x0〉;〈x7,x3〉;〈x8,x9〉;〈x2,x3〉;〈x0,x8〉;〈x7,xA〉;〈xA,xC〉;〈x5,xD〉;〈x6,xD〉;〈xC,xE〉;〈x0,xC〉;〈x1,xB〉;〈x1,x7〉;〈xC,x1〉;\r
+〈x4,x2〉;〈x5,x3〉;〈x1,x5〉;〈x7,xC〉;〈x7,x4〉;〈x2,xB〉;〈x2,x5〉;〈x5,x6〉;〈x6,x1〉;〈xE,xC〉;〈x0,xB〉;〈x4,x2〉;〈x0,x4〉;〈xC,xA〉;〈x0,x9〉;〈xA,xB〉;\r
+〈x1,xB〉;〈xD,x0〉;〈x9,xF〉;〈x6,xA〉;〈x7,xF〉;〈x4,x1〉;〈xF,x8〉;〈xE,xA〉;〈x8,x2〉;〈x8,x1〉;〈x4,x1〉;〈xC,xE〉;〈xC,xE〉;〈x0,xD〉;〈x2,xB〉;〈x3,x3〉;\r
+〈xA,x3〉;〈x6,x4〉;〈xF,xA〉;〈xA,x6〉;〈x3,x9〉;〈x7,xF〉;〈xF,x6〉;〈xB,x2〉;〈x5,x5〉;〈x6,xB〉;〈xA,xC〉;〈x3,x3〉;〈x9,x3〉;〈xE,x7〉;〈xB,xE〉;〈x3,x4〉;\r
+〈xC,xF〉;〈xE,xF〉;〈xA,x2〉;〈xE,xE〉;〈xE,xD〉;〈xC,xB〉;〈xB,x0〉;〈x8,x9〉;〈xD,xA〉;〈x3,xB〉;〈xB,xE〉;〈x3,xE〉;〈x3,x3〉;〈x5,x1〉;〈xA,x5〉;〈x3,xC〉;\r
+〈xC,xC〉;〈xA,x0〉;〈xF,xD〉;〈x3,x9〉;〈xC,xB〉;〈xF,xC〉;〈x1,xF〉;〈x8,xD〉;〈x6,x8〉;〈xD,x4〉;〈x8,xC〉;〈xA,xA〉;〈x8,xE〉;〈x3,xA〉;〈x9,x7〉;〈x2,x6〉;\r
+〈x6,xB〉;〈xA,xC〉;〈x8,xA〉;〈x4,xB〉;〈x7,x4〉;〈x3,xF〉;〈xB,x7〉;〈xB,xF〉;〈x0,xC〉;〈xE,x6〉;〈xC,xD〉;〈x4,x2〉;〈xF,xA〉;〈xE,xE〉;〈xF,x9〉;〈x0,xC〉;\r
+〈x2,xC〉;〈x7,x9〉;〈x7,xE〉;〈xD,x8〉;〈x4,x0〉;〈x7,xC〉;〈x3,x8〉;〈x4,x9〉;〈x7,x1〉;〈x7,x5〉;〈xB,x7〉;〈x3,x6〉;〈x0,x7〉;〈x1,xA〉;〈x2,xC〉;〈x1,xE〉;\r
+\r
+〈x3,xC〉;〈x7,xA〉;〈x3,x8〉;〈x4,xA〉;〈x3,x4〉;〈x2,x0〉;〈x9,x5〉;〈x6,x0〉;〈xF,x7〉;〈xC,x3〉;〈xB,x1〉;〈x6,xE〉;〈xB,x1〉;〈x7,x0〉;〈x7,x4〉;〈x3,xB〉;\r
+〈x0,xD〉;〈x6,xD〉;〈xF,xB〉;〈xE,x5〉;〈xE,x2〉;〈x6,x6〉;〈x6,x8〉;〈x0,x8〉;〈xF,xB〉;〈x3,xC〉;〈x8,xC〉;〈xD,xD〉;〈x0,x2〉;〈x2,xE〉;〈x6,xE〉;〈xF,x1〉;\r
+〈xA,xF〉;〈x7,x9〉;〈x6,x7〉;〈xE,x2〉;〈x4,xC〉;〈xA,x5〉;〈x7,x9〉;〈xC,x6〉;〈xB,x5〉;〈xA,xF〉;〈x1,x5〉;〈xF,xE〉;〈xE,x2〉;〈x2,xB〉;〈xC,xA〉;〈xE,x6〉;\r
+〈x3,xE〉;〈x2,xC〉;〈x5,x8〉;〈x7,x2〉;〈xC,xE〉;〈x7,x1〉;〈x8,xC〉;〈xB,xE〉;〈x2,x0〉;〈x6,x6〉;〈x0,x7〉;〈x6,xF〉;〈xD,x1〉;〈x8,x2〉;〈x3,x1〉;〈xF,x3〉;\r
+〈x9,x5〉;〈x9,x1〉;〈x1,x2〉;〈xF,x3〉;〈x4,xF〉;〈x6,xC〉;〈xA,x6〉;〈x8,xE〉;〈xB,x2〉;〈x7,x8〉;〈xD,xE〉;〈x7,x9〉;〈xC,x5〉;〈x2,x2〉;〈xF,x3〉;〈x0,x7〉;\r
+〈xE,x7〉;〈x9,xE〉;〈x9,x2〉;〈x7,x3〉;〈x3,xC〉;〈xA,x1〉;〈xD,xA〉;〈x2,x1〉;〈x2,x3〉;〈x4,x5〉;〈xE,x5〉;〈x7,x4〉;〈x8,x4〉;〈xC,x2〉;〈x6,x8〉;〈x3,x5〉;\r
+〈x9,xA〉;〈xC,x8〉;〈x2,xE〉;〈x1,xD〉;〈xD,x1〉;〈xA,x6〉;〈xD,xF〉;〈x0,x6〉;〈x7,xF〉;〈x8,xC〉;〈x2,x8〉;〈xF,x6〉;〈xC,x3〉;〈xF,x8〉;〈x6,x2〉;〈xF,x9〉;\r
+〈x5,x7〉;〈x2,x7〉;〈xD,x1〉;〈xD,x3〉;〈x0,xB〉;〈xA,x3〉;〈x8,x7〉;〈x8,x3〉;〈xC,x9〉;〈x1,x4〉;〈xB,x4〉;〈xC,x5〉;〈xE,xD〉;〈x5,x4〉;〈xE,x1〉;〈xB,x9〉;\r
+〈x2,x9〉;〈x6,xF〉;〈xE,x3〉;〈x0,xD〉;〈xC,xC〉;〈xF,x6〉;〈x0,xD〉;〈x2,x4〉;〈x4,x5〉;〈x6,xE〉;〈xE,x4〉;〈xD,xB〉;〈xF,x9〉;〈xC,x1〉;〈xD,x4〉;〈xB,x4〉;\r
+〈xB,x5〉;〈x6,x6〉;〈x1,xC〉;〈x6,xE〉;〈xA,xF〉;〈x4,x0〉;〈xE,x6〉;〈x4,x9〉;〈x7,xE〉;〈x4,x1〉;〈x1,x8〉;〈x8,x7〉;〈xD,xF〉;〈xB,xB〉;〈x6,x0〉;〈x0,x5〉;\r
+〈xF,x4〉;〈x5,xD〉;〈xA,x1〉;〈xD,x6〉;〈x6,x7〉;〈x2,xA〉;〈xC,x8〉;〈x7,x7〉;〈xF,x9〉;〈x8,xA〉;〈xF,x9〉;〈x2,x6〉;〈xE,xF〉;〈x7,x4〉;〈x5,x8〉;〈x6,xA〉;\r
+〈xC,x8〉;〈x3,x5〉;〈x1,x0〉;〈xC,x5〉;〈x1,xE〉;〈x0,xB〉;〈x8,x3〉;〈x6,xA〉;〈x4,x4〉;〈x8,xD〉;〈x5,xC〉;〈xF,xB〉;〈xF,xE〉;〈x9,x2〉;〈x0,x3〉;〈x4,x3〉;\r
+〈x0,x5〉;〈xA,xF〉;〈xC,x0〉;〈xF,x3〉;〈x0,x4〉;〈x2,x8〉;〈x9,xD〉;〈x0,x9〉;〈x3,x5〉;〈xE,x3〉;〈x5,xF〉;〈x4,x5〉;〈xA,xB〉;〈xC,xD〉;〈x8,xC〉;〈xF,xD〉;\r
+〈x2,xC〉;〈x9,xD〉;〈xA,xF〉;〈x6,x4〉;〈x4,x3〉;〈x8,x0〉;〈x8,x2〉;〈xE,x5〉;〈x8,xE〉;〈x3,xD〉;〈x2,xD〉;〈xD,xB〉;〈xD,xF〉;〈xA,xB〉;〈x0,x8〉;〈x1,x6〉;\r
+〈xE,xC〉;〈x7,xE〉;〈xA,x7〉;〈xC,xB〉;〈xD,x8〉;〈x5,xC〉;〈x2,xC〉;〈x8,x8〉;〈x9,x8〉;〈xC,x2〉;〈xA,xD〉;〈x1,xD〉;〈xB,x0〉;〈xB,x1〉;〈xC,xE〉;〈x9,x3〉;\r
+〈xE,x2〉;〈xF,x4〉;〈xD,xB〉;〈xA,x5〉;〈xB,x6〉;〈x4,x9〉;〈x8,x7〉;〈x1,xD〉;〈xA,x2〉;〈x7,x9〉;〈x3,x5〉;〈xB,xE〉;〈x5,x5〉;〈xC,xD〉;〈x6,x3〉;〈x2,xC〉;\r
+\r
+〈x1,x0〉;〈x1,xF〉;〈xE,xC〉;〈x3,xB〉;〈x8,xA〉;〈x3,xF〉;〈x3,x8〉;〈x8,x0〉;〈x1,xC〉;〈x2,xD〉;〈x9,x2〉;〈x5,xF〉;〈xE,x1〉;〈xB,x5〉;〈xB,xC〉;〈x8,x3〉;\r
+〈xB,x6〉;〈x1,xB〉;〈xE,xD〉;〈x4,xF〉;〈x3,xA〉;〈xC,x4〉;〈xF,xE〉;〈xF,xF〉;〈xC,xB〉;〈x8,x1〉;〈x6,x7〉;〈xC,x2〉;〈x5,x9〉;〈xD,xA〉;〈x0,xA〉;〈x9,xC〉;\r
+〈x2,x2〉;〈xE,xB〉;〈x9,x3〉;〈xE,x2〉;〈x7,xF〉;〈xA,xC〉;〈x4,xA〉;〈x8,x2〉;〈x8,x1〉;〈x3,xF〉;〈xE,xB〉;〈x8,xB〉;〈x0,xF〉;〈x9,xC〉;〈x4,x1〉;〈x8,x2〉;\r
+〈x9,xB〉;〈x7,xC〉;〈x5,x1〉;〈xA,x7〉;〈xA,xB〉;〈xA,xD〉;〈x9,x2〉;〈x1,x9〉;〈xF,x0〉;〈xF,xD〉;〈x9,x3〉;〈xF,x6〉;〈xA,xD〉;〈x2,x4〉;〈xC,xB〉;〈xD,xE〉;\r
+〈xB,x5〉;〈xA,xB〉;〈x8,x1〉;〈x5,x4〉;〈xA,xE〉;〈x2,x4〉;〈x6,x4〉;〈xD,x2〉;〈xD,x0〉;〈xF,xE〉;〈x3,x3〉;〈x2,xA〉;〈x7,x5〉;〈x0,x7〉;〈x8,xF〉;〈x3,xA〉;\r
+〈x1,x2〉;〈x9,xF〉;〈xB,xE〉;〈x1,xB〉;〈x1,xB〉;〈x1,xF〉;〈xC,x7〉;〈xF,x1〉;〈x7,xC〉;〈x9,x1〉;〈x5,xD〉;〈x3,x2〉;〈xD,x9〉;〈xD,x6〉;〈xE,xA〉;〈x0,x6〉;\r
+〈x5,xB〉;〈x6,x9〉;〈x6,xB〉;〈xA,xC〉;〈x0,x9〉;〈x1,x6〉;〈xA,x3〉;〈xC,xA〉;〈x8,xE〉;〈x8,x3〉;〈x3,x4〉;〈xB,x7〉;〈x4,x1〉;〈x2,x7〉;〈xB,x3〉;〈x0,x1〉;\r
+〈xE,x6〉;〈x0,x6〉;〈x8,xC〉;〈x0,x4〉;〈x3,xD〉;〈xB,xE〉;〈x2,xC〉;〈x6,x6〉;〈xB,x5〉;〈x8,x6〉;〈x1,x1〉;〈x1,x3〉;〈x6,xD〉;〈xD,x0〉;〈xB,xE〉;〈x8,xD〉;\r
+〈xC,x7〉;〈x5,x5〉;〈x0,x2〉;〈xC,x1〉;〈x7,x6〉;〈x5,xF〉;〈x2,x0〉;〈x5,xE〉;〈xE,x4〉;〈x3,xE〉;〈x7,x7〉;〈xE,x1〉;〈x3,xF〉;〈xE,x8〉;〈x6,xC〉;〈x4,xA〉;\r
+〈xA,x0〉;〈xF,xE〉;〈xC,xE〉;〈x3,xF〉;〈x6,x7〉;〈x9,x4〉;〈x3,xF〉;〈xE,xF〉;〈xE,xF〉;〈x8,x6〉;〈xD,x9〉;〈x4,xA〉;〈x0,x8〉;〈x8,xB〉;〈xC,x8〉;〈x1,xC〉;\r
+〈xA,xD〉;〈x2,x0〉;〈xA,x7〉;〈x8,xC〉;〈x0,x6〉;〈x6,x7〉;〈xA,xF〉;〈x7,x3〉;〈xC,xD〉;〈x1,x6〉;〈x8,x4〉;〈x3,x2〉;〈xD,x0〉;〈xF,x3〉;〈xD,xC〉;〈xD,xB〉;\r
+〈xB,x7〉;〈x2,x4〉;〈x6,xA〉;〈x6,x3〉;〈x1,xC〉;〈xA,x1〉;〈xD,xE〉;〈xB,xC〉;〈x9,x2〉;〈xF,x1〉;〈x5,xC〉;〈xE,x7〉;〈xE,x0〉;〈xD,x5〉;〈xA,x4〉;〈x4,xA〉;\r
+〈x0,x0〉;〈xD,x6〉;〈x2,x2〉;〈x9,xC〉;〈x5,x2〉;〈x8,xF〉;〈xE,x8〉;〈x2,x2〉;〈xA,x2〉;〈xF,x0〉;〈x9,x8〉;〈x3,x8〉;〈x0,xD〉;〈xF,x6〉;〈x4,x3〉;〈x7,x9〉;\r
+〈x8,x2〉;〈xA,xF〉;〈xD,x5〉;〈xC,x1〉;〈x8,x2〉;〈x5,x2〉;〈xD,xB〉;〈x8,xF〉;〈x7,xE〉;〈xD,x1〉;〈x9,xD〉;〈xA,x6〉;〈x8,xE〉;〈x9,xE〉;〈xA,x9〉;〈x8,xE〉;\r
+〈xD,x1〉;〈xF,x6〉;〈xB,x0〉;〈xE,xA〉;〈x8,x3〉;〈xE,x9〉;〈xF,x7〉;〈x3,xB〉;〈x4,xA〉;〈x0,x9〉;〈x1,xE〉;〈x3,x2〉;〈xD,x2〉;〈x5,xD〉;〈xD,x7〉;〈xA,xB〉;\r
+〈x4,xD〉;〈x6,xF〉;〈x5,x9〉;〈xF,xC〉;〈x4,x3〉;〈x4,x1〉;〈x0,x0〉;〈x3,xC〉;〈x9,x4〉;〈x5,x2〉;〈x5,x9〉;〈x6,xC〉;〈x6,xE〉;〈xE,x8〉;〈x6,x6〉;〈xF,x5〉;\r
+\r
+〈x9,xC〉;〈x5,x7〉;〈x6,xC〉;〈xE,x2〉;〈x3,xB〉;〈xA,x2〉;〈x2,x1〉;〈xE,xE〉;〈xF,x6〉;〈x4,xF〉;〈xF,x3〉;〈x6,x2〉;〈xD,xB〉;〈x8,xF〉;〈x6,x4〉;〈xD,x3〉;\r
+〈x8,x0〉;〈x6,x9〉;〈x9,x7〉;〈x4,x7〉;〈x8,xB〉;〈xB,x6〉;〈x3,x8〉;〈x4,x5〉;〈xB,xE〉;〈x0,xD〉;〈x6,x1〉;〈xC,xF〉;〈x7,x8〉;〈xC,xF〉;〈x4,x1〉;〈x7,xF〉;\r
+〈xF,x4〉;〈x5,xA〉;〈x8,xB〉;〈x7,x2〉;〈xE,x6〉;〈x7,xD〉;〈x4,xC〉;〈x1,x8〉;〈xE,xE〉;〈x3,xA〉;〈x1,x2〉;〈x9,x4〉;〈x3,x4〉;〈x3,x0〉;〈x3,x9〉;〈x0,x0〉;\r
+〈x9,x5〉;〈x6,x0〉;〈xF,xA〉;〈x7,xF〉;〈xA,x6〉;〈xC,x7〉;〈xB,x1〉;〈x7,xE〉;〈x0,xD〉;〈x2,x4〉;〈xF,xF〉;〈x4,x3〉;〈x7,x8〉;〈x8,x8〉;〈x6,xC〉;〈x0,x7〉;\r
+〈x7,x3〉;〈x9,x2〉;〈xC,x8〉;〈x0,xB〉;〈x5,x0〉;〈x9,x7〉;〈xF,x4〉;〈x1,xB〉;〈xD,xB〉;〈x4,x5〉;〈x6,x2〉;〈x9,x1〉;〈x8,x7〉;〈x5,xD〉;〈xF,x5〉;〈x6,x1〉;\r
+〈x3,x8〉;〈xF,x3〉;〈x8,xA〉;〈x4,xF〉;〈xD,xB〉;〈x3,xD〉;〈x4,x3〉;〈x2,xF〉;〈xB,xA〉;〈xA,x9〉;〈xB,x4〉;〈xA,x9〉;〈x2,x6〉;〈x7,xE〉;〈x8,xC〉;〈x1,x6〉;\r
+〈xF,xC〉;〈x1,xC〉;〈xA,x7〉;〈x9,xD〉;〈x7,xC〉;〈x8,x5〉;〈xA,x7〉;〈x3,x5〉;〈x8,x6〉;〈x4,xF〉;〈xC,x8〉;〈x9,xA〉;〈x0,xD〉;〈x2,x4〉;〈x0,x4〉;〈x5,x8〉;\r
+〈x4,x2〉;〈x3,x0〉;〈x7,x5〉;〈x8,xD〉;〈xB,xC〉;〈x7,x2〉;〈x3,x4〉;〈xF,x9〉;〈x4,x6〉;〈x5,xE〉;〈xF,x4〉;〈x5,xC〉;〈x6,x7〉;〈x0,x4〉;〈x8,x8〉;〈x2,x8〉;\r
+〈xA,x9〉;〈xE,xC〉;〈xE,xC〉;〈x7,xC〉;〈x4,x6〉;〈x6,xE〉;〈x1,xC〉;〈xD,x9〉;〈xC,x0〉;〈x6,x1〉;〈x5,x8〉;〈x5,xE〉;〈x3,x7〉;〈xB,x0〉;〈x9,x9〉;〈x9,x7〉;\r
+〈xA,x0〉;〈x9,xB〉;〈x7,x7〉;〈x4,xA〉;〈xD,x0〉;〈xB,x5〉;〈x1,xD〉;〈xA,x6〉;〈x4,xA〉;〈xC,x5〉;〈x3,xA〉;〈x9,x4〉;〈xB,x4〉;〈x5,x0〉;〈x0,x6〉;〈xC,x0〉;\r
+〈xA,x8〉;〈x0,x2〉;〈x5,xF〉;〈x0,xE〉;〈x2,x1〉;〈x0,x3〉;〈xB,x1〉;〈x9,x6〉;〈x0,x2〉;〈x9,x7〉;〈x8,x0〉;〈x1,xA〉;〈x0,xB〉;〈x3,xD〉;〈x2,x1〉;〈x7,xF〉;\r
+〈x0,x3〉;〈x2,x9〉;〈x3,x2〉;〈x6,x6〉;〈xB,xF〉;〈x3,xB〉;〈x5,x7〉;〈x5,x3〉;〈xE,x7〉;〈xD,x5〉;〈xE,x5〉;〈x4,x5〉;〈xA,x3〉;〈x1,xA〉;〈x1,xB〉;〈xF,x8〉;\r
+〈xC,xF〉;〈xB,x3〉;〈x9,x5〉;〈x5,x9〉;〈x4,xE〉;〈x6,x4〉;〈x4,x3〉;〈xF,x4〉;〈x2,x5〉;〈xC,xC〉;〈x6,x1〉;〈x3,x5〉;〈xD,xF〉;〈x3,x6〉;〈x5,x5〉;〈xC,xF〉;\r
+〈x9,xA〉;〈x1,x1〉;〈xF,x6〉;〈xD,x4〉;〈x4,xF〉;〈x9,xB〉;〈xA,xF〉;〈xF,x2〉;〈x0,x3〉;〈x1,x9〉;〈x9,xB〉;〈xA,xB〉;〈xC,x4〉;〈x1,x9〉;〈xA,x1〉;〈xE,xA〉;\r
+〈x1,xB〉;〈x2,x5〉;〈xA,xD〉;〈xA,xA〉;〈x0,x0〉;〈x5,xB〉;〈x9,xD〉;〈x6,xF〉;〈x8,x8〉;〈x6,xF〉;〈x3,x0〉;〈x8,x5〉;〈xC,x6〉;〈x1,x7〉;〈x5,x7〉;〈x1,x1〉;\r
+〈xA,xB〉;〈x0,x2〉;〈xD,xD〉;〈x9,x2〉;〈x4,xD〉;〈x8,x2〉;〈x0,x2〉;〈x3,x5〉;〈xC,xB〉;〈x4,x4〉;〈xA,x4〉;〈x4,x1〉;〈xD,x5〉;〈x1,x2〉;〈xE,x7〉;〈x4,xD〉;\r
diff --git a/helm/software/matita/library/freescale/doc/ordine_compilazione.txt b/helm/software/matita/library/freescale/doc/ordine_compilazione.txt
new file mode 100644 (file)
index 0000000..604286e
--- /dev/null
@@ -0,0 +1,40 @@
+\r
+tutte le definizioni fanno capo a\r
+ cic:/matita/freescale/...\r
+\r
+extra.ma\r
+exadecim.ma\r
+byte8.ma\r
+word16.ma\r
+aux_bases.ma\r
+opcode.ma\r
+table_HC05.ma\r
+table_HC08.ma\r
+table_HCS08.ma\r
+table_RS08.ma\r
+translation.ma\r
+memory_struct.ma\r
+memory_func.ma\r
+memory_trees.ma\r
+memory_bits.ma\r
+memory_abs.ma\r
+status.ma\r
+model.ma\r
+load_write.ma\r
+multivm.ma\r
+micro_tests.ma\r
+medium_tests_tools.ma\r
+medium_tests.ma\r
+tests.ma\r
+\r
+dubbi:\r
+1) atomicita' istruzione\r
+   - overflow su pc durante fetch opcode/immediati/indirizzi (ammesso)\r
+   - overflow su indiretto durante caricamento operandi (non ammesso)\r
+2) atomicita' salto\r
+   - overflow su pc in seguito a branch (ammesso)\r
+3) cosa succede con un ILLEGAL OPCODE\r
+   - nell'HC05\r
+   - nell'HC08/HCS08/RS08 con mascheramento ILLEGAL OPCODE\r
+4) cosa succede con un ILLEGAL ADDRESS\r
+   - nell'HC08/HCS08/RS08 con mascheramento ILLEGAL ADDRESS\r
diff --git a/helm/software/matita/library/freescale/doc/reverse.txt b/helm/software/matita/library/freescale/doc/reverse.txt
new file mode 100644 (file)
index 0000000..cf8cf32
--- /dev/null
@@ -0,0 +1,99 @@
+\r
+A: 0x00 HX: 0x0D4B SP: 0x0D4A PC: 0x18E0 Z:1\r
+A: 0x00 HX: LUNG/2 SP: 0x0D4A PC: 0x192B Z:1\r
+\r
+   8: 178806 -> 179480 :   674 = 42+79*   8\r
+  16: 178806 -> 180112 :  1306 = 42+79*  16\r
+  32: 178806 -> 181376 :  2570 = 42+79*  32\r
+  64: 178806 -> 183904 :  5098 = 42+79*  64\r
+ 128: 178806 -> 188960 : 10154 = 42+79* 128\r
+ 256: 178806 -> 199072 : 20266 = 42+79* 256\r
+ 511: 178806 -> 219138 : 40332 = 42+79* 511\r
+ 512: 178806 -> 219301 : 40495 < 42+79* 512 = 40490\r
+ 514: 178806 -> 219459 : 40653 < 42+79* 514 = 40648\r
+1024: 178806 -> 259754 : 80948 > 42+79*1024 = 80938\r
+\r
+ottima definizione esatta del tempo di esecuzione!\r
+\r
+T(n)=42+79*n+5*(n/512) oppure T(n)=42+79*n+(n>>9)\r
+\r
+dati 0x100-> 3072 byte\r
+\r
+**********************************************\r
+\r
+static unsigned char dati[3072]={...};\r
+\r
+void swap(unsigned char *a, unsigned char *b)\r
+ { unsigned char tmp=*a; *a=*b; *b=tmp; return; }\r
+\r
+18BE       PSHX\r
+18BF       PSHH\r
+18C0       LDHX 5,SP\r
+18C3       LDA ,X\r
+18C4       LDHX 1,SP\r
+18C7       PSHA\r
+18C8       LDA ,X\r
+18C9       LDHX 6,SP\r
+18CC       STA ,X\r
+18CD       LDHX 2,SP\r
+18D0       PULA\r
+18D1       STA ,X\r
+18D2       AIS #2\r
+18D4       RTS\r
+\r
+18D5-18DF ...\r
+\r
+void main(void)\r
+{\r
+ unsigned int pos=0,limit=0;\r
+\r
+ for(limit=3072;pos<(limit/2);pos++)\r
+  { swap(&dati[pos],&dati[limit-pos-1]); }\r
+\r
+18E0       LDHX #LUNG\r
+18E3       STHX 4,SP\r
+18E6 20 32 BRA *+52 ; 191A\r
+18E8       TSX\r
+18E9       LDA 2,X\r
+18EB       ADD #0x00\r
+18ED       PSHA\r
+18EE       LDA 1,X\r
+18F0       ADC #0x01\r
+18F2       PSHA\r
+18F3       LDA 4,X\r
+18F5       SUB 2,X\r
+18F7       STA ,X\r
+18F8       LDA 3,X\r
+18FA       SBC 1,X\r
+18FC       PSHA\r
+18FD       LDX ,X\r
+18FE       PULH\r
+18FF       AIX #-1\r
+1901       TXA\r
+1902       ADD #0x00\r
+1904       PSHH\r
+1905       TSX\r
+1906       STA 3,X\r
+1908       PULA\r
+1909       ADC #0x01\r
+190B       LDX 3,X\r
+190D       PSHA\r
+190E       PULH\r
+190F AD AD BSR *-81 ; 18BE\r
+1911       AIS #2\r
+1913       TSX\r
+1914       INC 2,X\r
+1916 26 02 BNE *+4 ; 191A\r
+1918       INC 1,X\r
+191A       TSX\r
+191B       LDA 3,X\r
+191D       PSHA\r
+191E       PULH\r
+191F       LSRA\r
+1920       TSX\r
+1921       LDX 4,X\r
+1923       RORX\r
+1924       PSHA\r
+1925       PULH\r
+1926       CPHX 2,SP\r
+1929 22 BD BHI *-65 ; 18E8\r
diff --git a/helm/software/matita/library/freescale/exadecim.ma b/helm/software/matita/library/freescale/exadecim.ma
new file mode 100644 (file)
index 0000000..21a31c9
--- /dev/null
@@ -0,0 +1,1826 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/exadecim/".
+
+(*include "/media/VIRTUOSO/freescale/extra.ma".*)
+include "freescale/extra.ma".
+
+(* ***************************** *)
+(* DEFINIZIONE DEGLI ESADECIMALI *)
+(* ***************************** *)
+
+inductive exadecim : Type ≝
+  x0: exadecim
+| x1: exadecim
+| x2: exadecim
+| x3: exadecim
+| x4: exadecim
+| x5: exadecim
+| x6: exadecim
+| x7: exadecim
+| x8: exadecim
+| x9: exadecim
+| xA: exadecim
+| xB: exadecim
+| xC: exadecim
+| xD: exadecim
+| xE: exadecim
+| xF: exadecim.
+
+(* operatore = *)
+definition eq_ex ≝
+λe1,e2:exadecim.
+ match e1 with
+  [ x0 ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x1 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ true  | x2 ⇒ false | x3 ⇒ false
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x2 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ true  | x3 ⇒ false
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x3 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ true
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x4 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
+   | x4 ⇒ true  | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x5 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
+   | x4 ⇒ false | x5 ⇒ true  | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x6 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ true  | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x7 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ true 
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x8 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ true  | x9 ⇒ false | xA ⇒ false | xB ⇒ false
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x9 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ true  | xA ⇒ false | xB ⇒ false
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | xA ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ true  | xB ⇒ false
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | xB ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ true 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | xC ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
+   | xC ⇒ true  | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | xD ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
+   | xC ⇒ false | xD ⇒ true  | xE ⇒ false | xF ⇒ false ] 
+  | xE ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
+   | xC ⇒ false | xD ⇒ false | xE ⇒ true  | xF ⇒ false ] 
+  | xF ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ true ]
+  ].
+
+(* operatore < *)
+definition lt_ex ≝
+λe1,e2:exadecim.
+ match e1 with
+  [ x0 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ true | x2 ⇒ true | x3 ⇒ true
+   | x4 ⇒ true  | x5 ⇒ true | x6 ⇒ true | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true | xA ⇒ true | xB ⇒ true
+   | xC ⇒ true  | xD ⇒ true | xE ⇒ true | xF ⇒ true ] 
+  | x1 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ true | x3 ⇒ true
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true | xB ⇒ true
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true | xF ⇒ true ] 
+  | x2 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ true
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ] 
+  | x3 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ] 
+  | x4 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
+   | x4 ⇒ false | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ] 
+  | x5 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true 
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ] 
+  | x6 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true 
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ] 
+  | x7 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false 
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ] 
+  | x8 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true 
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ] 
+  | x9 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ true  | xB ⇒ true 
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ] 
+  | xA ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ true
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ] 
+  | xB ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ] 
+  | xC ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ] 
+  | xD ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ true  | xF ⇒ true ] 
+  | xE ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ true ] 
+  | xF ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ]
+  ].
+
+(* operatore ≤ *)
+definition le_ex ≝
+λe1,e2:exadecim.
+ match e1 with
+  [ x0 ⇒ match e2 with
+   [ x0 ⇒ true | x1 ⇒ true | x2 ⇒ true | x3 ⇒ true 
+   | x4 ⇒ true | x5 ⇒ true | x6 ⇒ true | x7 ⇒ true
+   | x8 ⇒ true | x9 ⇒ true | xA ⇒ true | xB ⇒ true 
+   | xC ⇒ true | xD ⇒ true | xE ⇒ true | xF ⇒ true ] 
+  | x1 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ true | x2 ⇒ true | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true | x6 ⇒ true | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true | xA ⇒ true | xB ⇒ true 
+   | xC ⇒ true  | xD ⇒ true | xE ⇒ true | xF ⇒ true ] 
+  | x2 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ true | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true | xB ⇒ true 
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true | xF ⇒ true ] 
+  | x3 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true 
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ] 
+  | x4 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true 
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ] 
+  | x5 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true 
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ] 
+  | x6 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true 
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ] 
+  | x7 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ true 
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true 
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ] 
+  | x8 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true 
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ] 
+  | x9 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true 
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ] 
+  | xA ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ true  | xB ⇒ true 
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ] 
+  | xB ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ true 
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ] 
+  | xC ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ] 
+  | xD ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ] 
+  | xE ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ true  | xF ⇒ true ] 
+  | xF ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ true ]
+  ].
+
+(* operatore > *)
+definition gt_ex ≝
+λe1,e2:exadecim.
+ match e1 with
+  [ x0 ⇒ match e2 with
+   [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x1 ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x2 ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x3 ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x4 ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x5 ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x6 ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x7 ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ]
+  | x8 ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ]
+  | x9 ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ]
+  | xA ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ]
+  | xB ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ]
+  | xC ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ]
+  | xD ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true 
+   | xC ⇒ true  | xD ⇒ false | xE ⇒ false | xF ⇒ false ]
+  | xE ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true 
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ false | xF ⇒ false ]
+  | xF ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true 
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ false ]
+  ].
+
+(* operatore ≥ *)
+definition ge_ex ≝
+λe1,e2:exadecim.
+ match e1 with
+  [ x0 ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x1 ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ false | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x2 ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ false 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x3 ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x4 ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x5 ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ false | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x6 ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ false
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ] 
+  | x7 ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ false | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ]
+  | x8 ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ false | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ]
+  | x9 ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ false | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ]
+  | xA ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ false 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ]
+  | xB ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true 
+   | xC ⇒ false | xD ⇒ false | xE ⇒ false | xF ⇒ false ]
+  | xC ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true 
+   | xC ⇒ true  | xD ⇒ false | xE ⇒ false | xF ⇒ false ]
+  | xD ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true 
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ false | xF ⇒ false ]
+  | xE ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true 
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ false ]
+  | xF ⇒ match e2 with
+   [ x0 ⇒ true  | x1 ⇒ true  | x2 ⇒ true  | x3 ⇒ true 
+   | x4 ⇒ true  | x5 ⇒ true  | x6 ⇒ true  | x7 ⇒ true
+   | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true 
+   | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ]
+  ].
+
+(* operatore and *)
+definition and_ex ≝
+λe1,e2:exadecim.match e1 with
+ [ x0 ⇒ match e2 with
+  [ x0 ⇒ x0 | x1 ⇒ x0 | x2 ⇒ x0 | x3 ⇒ x0 
+  | x4 ⇒ x0 | x5 ⇒ x0 | x6 ⇒ x0 | x7 ⇒ x0
+  | x8 ⇒ x0 | x9 ⇒ x0 | xA ⇒ x0 | xB ⇒ x0 
+  | xC ⇒ x0 | xD ⇒ x0 | xE ⇒ x0 | xF ⇒ x0 ]
+ | x1 ⇒ match e2 with
+  [ x0 ⇒ x0 | x1 ⇒ x1 | x2 ⇒ x0 | x3 ⇒ x1 
+  | x4 ⇒ x0 | x5 ⇒ x1 | x6 ⇒ x0 | x7 ⇒ x1
+  | x8 ⇒ x0 | x9 ⇒ x1 | xA ⇒ x0 | xB ⇒ x1 
+  | xC ⇒ x0 | xD ⇒ x1 | xE ⇒ x0 | xF ⇒ x1 ]
+ | x2 ⇒ match e2 with
+  [ x0 ⇒ x0 | x1 ⇒ x0 | x2 ⇒ x2 | x3 ⇒ x2 
+  | x4 ⇒ x0 | x5 ⇒ x0 | x6 ⇒ x2 | x7 ⇒ x2
+  | x8 ⇒ x0 | x9 ⇒ x0 | xA ⇒ x2 | xB ⇒ x2 
+  | xC ⇒ x0 | xD ⇒ x0 | xE ⇒ x2 | xF ⇒ x2 ]
+ | x3 ⇒ match e2 with
+  [ x0 ⇒ x0 | x1 ⇒ x1 | x2 ⇒ x2 | x3 ⇒ x3 
+  | x4 ⇒ x0 | x5 ⇒ x1 | x6 ⇒ x2 | x7 ⇒ x3
+  | x8 ⇒ x0 | x9 ⇒ x1 | xA ⇒ x2 | xB ⇒ x3 
+  | xC ⇒ x0 | xD ⇒ x1 | xE ⇒ x2 | xF ⇒ x3 ]
+ | x4 ⇒ match e2 with
+  [ x0 ⇒ x0 | x1 ⇒ x0 | x2 ⇒ x0 | x3 ⇒ x0 
+  | x4 ⇒ x4 | x5 ⇒ x4 | x6 ⇒ x4 | x7 ⇒ x4
+  | x8 ⇒ x0 | x9 ⇒ x0 | xA ⇒ x0 | xB ⇒ x0 
+  | xC ⇒ x4 | xD ⇒ x4 | xE ⇒ x4 | xF ⇒ x4 ]
+ | x5 ⇒ match e2 with
+  [ x0 ⇒ x0 | x1 ⇒ x1 | x2 ⇒ x0 | x3 ⇒ x1 
+  | x4 ⇒ x4 | x5 ⇒ x5 | x6 ⇒ x4 | x7 ⇒ x5
+  | x8 ⇒ x0 | x9 ⇒ x1 | xA ⇒ x0 | xB ⇒ x1 
+  | xC ⇒ x4 | xD ⇒ x5 | xE ⇒ x4 | xF ⇒ x5 ]
+ | x6 ⇒ match e2 with
+  [ x0 ⇒ x0 | x1 ⇒ x0 | x2 ⇒ x2 | x3 ⇒ x2 
+  | x4 ⇒ x4 | x5 ⇒ x4 | x6 ⇒ x6 | x7 ⇒ x6
+  | x8 ⇒ x0 | x9 ⇒ x0 | xA ⇒ x2 | xB ⇒ x2 
+  | xC ⇒ x4 | xD ⇒ x4 | xE ⇒ x6 | xF ⇒ x6 ]
+ | x7 ⇒ match e2 with
+  [ x0 ⇒ x0 | x1 ⇒ x1 | x2 ⇒ x2 | x3 ⇒ x3 
+  | x4 ⇒ x4 | x5 ⇒ x5 | x6 ⇒ x6 | x7 ⇒ x7
+  | x8 ⇒ x0 | x9 ⇒ x1 | xA ⇒ x2 | xB ⇒ x3 
+  | xC ⇒ x4 | xD ⇒ x5 | xE ⇒ x6 | xF ⇒ x7 ]
+ | x8 ⇒ match e2 with
+  [ x0 ⇒ x0 | x1 ⇒ x0 | x2 ⇒ x0 | x3 ⇒ x0 
+  | x4 ⇒ x0 | x5 ⇒ x0 | x6 ⇒ x0 | x7 ⇒ x0
+  | x8 ⇒ x8 | x9 ⇒ x8 | xA ⇒ x8 | xB ⇒ x8 
+  | xC ⇒ x8 | xD ⇒ x8 | xE ⇒ x8 | xF ⇒ x8 ]
+ | x9 ⇒ match e2 with
+  [ x0 ⇒ x0 | x1 ⇒ x1 | x2 ⇒ x0 | x3 ⇒ x1 
+  | x4 ⇒ x0 | x5 ⇒ x1 | x6 ⇒ x0 | x7 ⇒ x1
+  | x8 ⇒ x8 | x9 ⇒ x9 | xA ⇒ x8 | xB ⇒ x9 
+  | xC ⇒ x8 | xD ⇒ x9 | xE ⇒ x8 | xF ⇒ x9 ]
+ | xA ⇒ match e2 with
+  [ x0 ⇒ x0 | x1 ⇒ x0 | x2 ⇒ x2 | x3 ⇒ x2 
+  | x4 ⇒ x0 | x5 ⇒ x0 | x6 ⇒ x2 | x7 ⇒ x2
+  | x8 ⇒ x8 | x9 ⇒ x8 | xA ⇒ xA | xB ⇒ xA 
+  | xC ⇒ x8 | xD ⇒ x8 | xE ⇒ xA | xF ⇒ xA ]
+ | xB ⇒ match e2 with
+  [ x0 ⇒ x0 | x1 ⇒ x1 | x2 ⇒ x2 | x3 ⇒ x3 
+  | x4 ⇒ x0 | x5 ⇒ x1 | x6 ⇒ x2 | x7 ⇒ x3
+  | x8 ⇒ x8 | x9 ⇒ x9 | xA ⇒ xA | xB ⇒ xB 
+  | xC ⇒ x8 | xD ⇒ x9 | xE ⇒ xA | xF ⇒ xB ]
+ | xC ⇒ match e2 with
+  [ x0 ⇒ x0 | x1 ⇒ x0 | x2 ⇒ x0 | x3 ⇒ x0 
+  | x4 ⇒ x4 | x5 ⇒ x4 | x6 ⇒ x4 | x7 ⇒ x4
+  | x8 ⇒ x8 | x9 ⇒ x8 | xA ⇒ x8 | xB ⇒ x8 
+  | xC ⇒ xC | xD ⇒ xC | xE ⇒ xC | xF ⇒ xC ] 
+ | xD ⇒ match e2 with
+  [ x0 ⇒ x0 | x1 ⇒ x1 | x2 ⇒ x0 | x3 ⇒ x1 
+  | x4 ⇒ x4 | x5 ⇒ x5 | x6 ⇒ x4 | x7 ⇒ x5
+  | x8 ⇒ x8 | x9 ⇒ x9 | xA ⇒ x8 | xB ⇒ x9 
+  | xC ⇒ xC | xD ⇒ xD | xE ⇒ xC | xF ⇒ xD ] 
+ | xE ⇒ match e2 with
+  [ x0 ⇒ x0 | x1 ⇒ x0 | x2 ⇒ x2 | x3 ⇒ x2 
+  | x4 ⇒ x4 | x5 ⇒ x4 | x6 ⇒ x6 | x7 ⇒ x6
+  | x8 ⇒ x8 | x9 ⇒ x8 | xA ⇒ xA | xB ⇒ xA 
+  | xC ⇒ xC | xD ⇒ xC | xE ⇒ xE | xF ⇒ xE ] 
+ | xF ⇒ match e2 with
+  [ x0 ⇒ x0 | x1 ⇒ x1 | x2 ⇒ x2 | x3 ⇒ x3 
+  | x4 ⇒ x4 | x5 ⇒ x5 | x6 ⇒ x6 | x7 ⇒ x7
+  | x8 ⇒ x8 | x9 ⇒ x9 | xA ⇒ xA | xB ⇒ xB 
+  | xC ⇒ xC | xD ⇒ xD | xE ⇒ xE | xF ⇒ xF ]
+ ]. 
+
+(* operatore or *)
+definition or_ex ≝
+λe1,e2:exadecim.match e1 with
+ [ x0 ⇒ match e2 with
+  [ x0 ⇒ x0 | x1 ⇒ x1 | x2 ⇒ x2 | x3 ⇒ x3 
+  | x4 ⇒ x4 | x5 ⇒ x5 | x6 ⇒ x6 | x7 ⇒ x7
+  | x8 ⇒ x8 | x9 ⇒ x9 | xA ⇒ xA | xB ⇒ xB 
+  | xC ⇒ xC | xD ⇒ xD | xE ⇒ xE | xF ⇒ xF ]
+ | x1 ⇒ match e2 with
+  [ x0 ⇒ x1 | x1 ⇒ x1 | x2 ⇒ x3 | x3 ⇒ x3 
+  | x4 ⇒ x5 | x5 ⇒ x5 | x6 ⇒ x7 | x7 ⇒ x7
+  | x8 ⇒ x9 | x9 ⇒ x9 | xA ⇒ xB | xB ⇒ xB 
+  | xC ⇒ xD | xD ⇒ xD | xE ⇒ xF | xF ⇒ xF ]
+ | x2 ⇒ match e2 with
+  [ x0 ⇒ x2 | x1 ⇒ x3 | x2 ⇒ x2 | x3 ⇒ x3 
+  | x4 ⇒ x6 | x5 ⇒ x7 | x6 ⇒ x6 | x7 ⇒ x7
+  | x8 ⇒ xA | x9 ⇒ xB | xA ⇒ xA | xB ⇒ xB 
+  | xC ⇒ xE | xD ⇒ xF | xE ⇒ xE | xF ⇒ xF ]
+ | x3 ⇒ match e2 with
+  [ x0 ⇒ x3 | x1 ⇒ x3 | x2 ⇒ x3 | x3 ⇒ x3 
+  | x4 ⇒ x7 | x5 ⇒ x7 | x6 ⇒ x7 | x7 ⇒ x7
+  | x8 ⇒ xB | x9 ⇒ xB | xA ⇒ xB | xB ⇒ xB 
+  | xC ⇒ xF | xD ⇒ xF | xE ⇒ xF | xF ⇒ xF ]
+ | x4 ⇒ match e2 with
+  [ x0 ⇒ x4 | x1 ⇒ x5 | x2 ⇒ x6 | x3 ⇒ x7 
+  | x4 ⇒ x4 | x5 ⇒ x5 | x6 ⇒ x6 | x7 ⇒ x7
+  | x8 ⇒ xC | x9 ⇒ xD | xA ⇒ xE | xB ⇒ xF 
+  | xC ⇒ xC | xD ⇒ xD | xE ⇒ xE | xF ⇒ xF ]
+ | x5 ⇒ match e2 with
+  [ x0 ⇒ x5 | x1 ⇒ x5 | x2 ⇒ x7 | x3 ⇒ x7 
+  | x4 ⇒ x5 | x5 ⇒ x5 | x6 ⇒ x7 | x7 ⇒ x7
+  | x8 ⇒ xD | x9 ⇒ xD | xA ⇒ xF | xB ⇒ xF 
+  | xC ⇒ xD | xD ⇒ xD | xE ⇒ xF | xF ⇒ xF ]
+ | x6 ⇒ match e2 with
+  [ x0 ⇒ x6 | x1 ⇒ x7 | x2 ⇒ x6 | x3 ⇒ x7 
+  | x4 ⇒ x6 | x5 ⇒ x7 | x6 ⇒ x6 | x7 ⇒ x7
+  | x8 ⇒ xE | x9 ⇒ xF | xA ⇒ xE | xB ⇒ xF 
+  | xC ⇒ xE | xD ⇒ xF | xE ⇒ xE | xF ⇒ xF ]
+ | x7 ⇒ match e2 with
+  [ x0 ⇒ x7 | x1 ⇒ x7 | x2 ⇒ x7 | x3 ⇒ x7 
+  | x4 ⇒ x7 | x5 ⇒ x7 | x6 ⇒ x7 | x7 ⇒ x7
+  | x8 ⇒ xF | x9 ⇒ xF | xA ⇒ xF | xB ⇒ xF 
+  | xC ⇒ xF | xD ⇒ xF | xE ⇒ xF | xF ⇒ xF ]
+ | x8 ⇒ match e2 with
+  [ x0 ⇒ x8 | x1 ⇒ x9 | x2 ⇒ xA | x3 ⇒ xB 
+  | x4 ⇒ xC | x5 ⇒ xD | x6 ⇒ xE | x7 ⇒ xF
+  | x8 ⇒ x8 | x9 ⇒ x9 | xA ⇒ xA | xB ⇒ xB 
+  | xC ⇒ xC | xD ⇒ xD | xE ⇒ xE | xF ⇒ xF ]
+ | x9 ⇒ match e2 with
+  [ x0 ⇒ x9 | x1 ⇒ x9 | x2 ⇒ xB | x3 ⇒ xB 
+  | x4 ⇒ xD | x5 ⇒ xD | x6 ⇒ xF | x7 ⇒ xF
+  | x8 ⇒ x9 | x9 ⇒ x9 | xA ⇒ xB | xB ⇒ xB 
+  | xC ⇒ xD | xD ⇒ xD | xE ⇒ xF | xF ⇒ xF ]
+ | xA ⇒ match e2 with
+  [ x0 ⇒ xA | x1 ⇒ xB | x2 ⇒ xA | x3 ⇒ xB 
+  | x4 ⇒ xE | x5 ⇒ xF | x6 ⇒ xE | x7 ⇒ xF
+  | x8 ⇒ xA | x9 ⇒ xB | xA ⇒ xA | xB ⇒ xB 
+  | xC ⇒ xE | xD ⇒ xF | xE ⇒ xE | xF ⇒ xF ]
+ | xB ⇒ match e2 with
+  [ x0 ⇒ xB | x1 ⇒ xB | x2 ⇒ xB | x3 ⇒ xB 
+  | x4 ⇒ xF | x5 ⇒ xF | x6 ⇒ xF | x7 ⇒ xF
+  | x8 ⇒ xB | x9 ⇒ xB | xA ⇒ xB | xB ⇒ xB 
+  | xC ⇒ xF | xD ⇒ xF | xE ⇒ xF | xF ⇒ xF ]
+ | xC ⇒ match e2 with
+  [ x0 ⇒ xC | x1 ⇒ xD | x2 ⇒ xE | x3 ⇒ xF 
+  | x4 ⇒ xC | x5 ⇒ xD | x6 ⇒ xE | x7 ⇒ xF
+  | x8 ⇒ xC | x9 ⇒ xD | xA ⇒ xE | xB ⇒ xF 
+  | xC ⇒ xC | xD ⇒ xD | xE ⇒ xE | xF ⇒ xF ]
+ | xD ⇒ match e2 with
+  [ x0 ⇒ xD | x1 ⇒ xD | x2 ⇒ xF | x3 ⇒ xF 
+  | x4 ⇒ xD | x5 ⇒ xD | x6 ⇒ xF | x7 ⇒ xF
+  | x8 ⇒ xD | x9 ⇒ xD | xA ⇒ xF | xB ⇒ xF 
+  | xC ⇒ xD | xD ⇒ xD | xE ⇒ xF | xF ⇒ xF ]
+ | xE ⇒ match e2 with
+  [ x0 ⇒ xE | x1 ⇒ xF | x2 ⇒ xE | x3 ⇒ xF 
+  | x4 ⇒ xE | x5 ⇒ xF | x6 ⇒ xE | x7 ⇒ xF
+  | x8 ⇒ xE | x9 ⇒ xF | xA ⇒ xE | xB ⇒ xF 
+  | xC ⇒ xE | xD ⇒ xF | xE ⇒ xE | xF ⇒ xF ]
+ | xF ⇒ match e2 with
+  [ x0 ⇒ xF | x1 ⇒ xF | x2 ⇒ xF | x3 ⇒ xF 
+  | x4 ⇒ xF | x5 ⇒ xF | x6 ⇒ xF | x7 ⇒ xF
+  | x8 ⇒ xF | x9 ⇒ xF | xA ⇒ xF | xB ⇒ xF 
+  | xC ⇒ xF | xD ⇒ xF | xE ⇒ xF | xF ⇒ xF ]
+ ].
+
+(* operatore xor *)
+definition xor_ex ≝
+λe1,e2:exadecim.match e1 with
+ [ x0 ⇒ match e2 with
+  [ x0 ⇒ x0 | x1 ⇒ x1 | x2 ⇒ x2 | x3 ⇒ x3 
+  | x4 ⇒ x4 | x5 ⇒ x5 | x6 ⇒ x6 | x7 ⇒ x7
+  | x8 ⇒ x8 | x9 ⇒ x9 | xA ⇒ xA | xB ⇒ xB 
+  | xC ⇒ xC | xD ⇒ xD | xE ⇒ xE | xF ⇒ xF ] 
+ | x1 ⇒ match e2 with
+  [ x0 ⇒ x1 | x1 ⇒ x0 | x2 ⇒ x3 | x3 ⇒ x2 
+  | x4 ⇒ x5 | x5 ⇒ x4 | x6 ⇒ x7 | x7 ⇒ x6
+  | x8 ⇒ x9 | x9 ⇒ x8 | xA ⇒ xB | xB ⇒ xA 
+  | xC ⇒ xD | xD ⇒ xC | xE ⇒ xF | xF ⇒ xE ] 
+ | x2 ⇒ match e2 with
+  [ x0 ⇒ x2 | x1 ⇒ x3 | x2 ⇒ x0 | x3 ⇒ x1 
+  | x4 ⇒ x6 | x5 ⇒ x7 | x6 ⇒ x4 | x7 ⇒ x5
+  | x8 ⇒ xA | x9 ⇒ xB | xA ⇒ x8 | xB ⇒ x9 
+  | xC ⇒ xE | xD ⇒ xF | xE ⇒ xC | xF ⇒ xD ] 
+ | x3 ⇒ match e2 with
+  [ x0 ⇒ x3 | x1 ⇒ x2 | x2 ⇒ x1 | x3 ⇒ x0 
+  | x4 ⇒ x7 | x5 ⇒ x6 | x6 ⇒ x5 | x7 ⇒ x4
+  | x8 ⇒ xB | x9 ⇒ xA | xA ⇒ x9 | xB ⇒ x8 
+  | xC ⇒ xF | xD ⇒ xE | xE ⇒ xD | xF ⇒ xC ] 
+ | x4 ⇒ match e2 with
+  [ x0 ⇒ x4 | x1 ⇒ x5 | x2 ⇒ x6 | x3 ⇒ x7 
+  | x4 ⇒ x0 | x5 ⇒ x1 | x6 ⇒ x2 | x7 ⇒ x3
+  | x8 ⇒ xC | x9 ⇒ xD | xA ⇒ xE | xB ⇒ xF 
+  | xC ⇒ x8 | xD ⇒ x9 | xE ⇒ xA | xF ⇒ xB ] 
+ | x5 ⇒ match e2 with
+  [ x0 ⇒ x5 | x1 ⇒ x4 | x2 ⇒ x7 | x3 ⇒ x6 
+  | x4 ⇒ x1 | x5 ⇒ x0 | x6 ⇒ x3 | x7 ⇒ x2
+  | x8 ⇒ xD | x9 ⇒ xC | xA ⇒ xF | xB ⇒ xE 
+  | xC ⇒ x9 | xD ⇒ x8 | xE ⇒ xB | xF ⇒ xA ] 
+ | x6 ⇒ match e2 with
+  [ x0 ⇒ x6 | x1 ⇒ x7 | x2 ⇒ x4 | x3 ⇒ x5 
+  | x4 ⇒ x2 | x5 ⇒ x3 | x6 ⇒ x0 | x7 ⇒ x1
+  | x8 ⇒ xE | x9 ⇒ xF | xA ⇒ xC | xB ⇒ xD 
+  | xC ⇒ xA | xD ⇒ xB | xE ⇒ x8 | xF ⇒ x9 ] 
+ | x7 ⇒ match e2 with
+  [ x0 ⇒ x7 | x1 ⇒ x6 | x2 ⇒ x5 | x3 ⇒ x4 
+  | x4 ⇒ x3 | x5 ⇒ x2 | x6 ⇒ x1 | x7 ⇒ x0
+  | x8 ⇒ xF | x9 ⇒ xE | xA ⇒ xD | xB ⇒ xC 
+  | xC ⇒ xB | xD ⇒ xA | xE ⇒ x9 | xF ⇒ x8 ] 
+ | x8 ⇒ match e2 with
+  [ x0 ⇒ x8 | x1 ⇒ x9 | x2 ⇒ xA | x3 ⇒ xB 
+  | x4 ⇒ xC | x5 ⇒ xD | x6 ⇒ xE | x7 ⇒ xF
+  | x8 ⇒ x0 | x9 ⇒ x1 | xA ⇒ x2 | xB ⇒ x3 
+  | xC ⇒ x4 | xD ⇒ x5 | xE ⇒ x6 | xF ⇒ x7 ] 
+ | x9 ⇒ match e2 with
+  [ x0 ⇒ x9 | x1 ⇒ x8 | x2 ⇒ xB | x3 ⇒ xA 
+  | x4 ⇒ xD | x5 ⇒ xC | x6 ⇒ xF | x7 ⇒ xE
+  | x8 ⇒ x1 | x9 ⇒ x0 | xA ⇒ x3 | xB ⇒ x2 
+  | xC ⇒ x5 | xD ⇒ x4 | xE ⇒ x7 | xF ⇒ x6 ] 
+ | xA ⇒ match e2 with
+  [ x0 ⇒ xA | x1 ⇒ xB | x2 ⇒ x8 | x3 ⇒ x9 
+  | x4 ⇒ xE | x5 ⇒ xF | x6 ⇒ xC | x7 ⇒ xD
+  | x8 ⇒ x2 | x9 ⇒ x3 | xA ⇒ x0 | xB ⇒ x1 
+  | xC ⇒ x6 | xD ⇒ x7 | xE ⇒ x4 | xF ⇒ x5 ] 
+ | xB ⇒ match e2 with
+  [ x0 ⇒ xB | x1 ⇒ xA | x2 ⇒ x9 | x3 ⇒ x8 
+  | x4 ⇒ xF | x5 ⇒ xE | x6 ⇒ xD | x7 ⇒ xC
+  | x8 ⇒ x3 | x9 ⇒ x2 | xA ⇒ x1 | xB ⇒ x0 
+  | xC ⇒ x7 | xD ⇒ x6 | xE ⇒ x5 | xF ⇒ x4 ]
+ | xC ⇒ match e2 with
+  [ x0 ⇒ xC | x1 ⇒ xD | x2 ⇒ xE | x3 ⇒ xF 
+  | x4 ⇒ x8 | x5 ⇒ x9 | x6 ⇒ xA | x7 ⇒ xB
+  | x8 ⇒ x4 | x9 ⇒ x5 | xA ⇒ x6 | xB ⇒ x7 
+  | xC ⇒ x0 | xD ⇒ x1 | xE ⇒ x2 | xF ⇒ x3 ] 
+ | xD ⇒ match e2 with
+  [ x0 ⇒ xD | x1 ⇒ xC | x2 ⇒ xF | x3 ⇒ xE 
+  | x4 ⇒ x9 | x5 ⇒ x8 | x6 ⇒ xB | x7 ⇒ xA
+  | x8 ⇒ x5 | x9 ⇒ x4 | xA ⇒ x7 | xB ⇒ x6 
+  | xC ⇒ x1 | xD ⇒ x0 | xE ⇒ x3 | xF ⇒ x2 ]
+ | xE ⇒ match e2 with
+  [ x0 ⇒ xE | x1 ⇒ xF | x2 ⇒ xC | x3 ⇒ xD 
+  | x4 ⇒ xA | x5 ⇒ xB | x6 ⇒ x8 | x7 ⇒ x9
+  | x8 ⇒ x6 | x9 ⇒ x7 | xA ⇒ x4 | xB ⇒ x5 
+  | xC ⇒ x2 | xD ⇒ x3 | xE ⇒ x0 | xF ⇒ x1 ] 
+ | xF ⇒ match e2 with
+  [ x0 ⇒ xF | x1 ⇒ xE | x2 ⇒ xD | x3 ⇒ xC 
+  | x4 ⇒ xB | x5 ⇒ xA | x6 ⇒ x9 | x7 ⇒ x8
+  | x8 ⇒ x7 | x9 ⇒ x6 | xA ⇒ x5 | xB ⇒ x4 
+  | xC ⇒ x3 | xD ⇒ x2 | xE ⇒ x1 | xF ⇒ x0 ] 
+ ].
+
+(* operatore rotazione destra con carry *)
+definition rcr_ex ≝
+λe:exadecim.λc:bool.match c with
+ [ true ⇒ match e with
+  [ x0 ⇒ pairT exadecim bool x8 false | x1 ⇒ pairT exadecim bool x8 true
+  | x2 ⇒ pairT exadecim bool x9 false | x3 ⇒ pairT exadecim bool x9 true
+  | x4 ⇒ pairT exadecim bool xA false | x5 ⇒ pairT exadecim bool xA true
+  | x6 ⇒ pairT exadecim bool xB false | x7 ⇒ pairT exadecim bool xB true
+  | x8 ⇒ pairT exadecim bool xC false | x9 ⇒ pairT exadecim bool xC true
+  | xA ⇒ pairT exadecim bool xD false | xB ⇒ pairT exadecim bool xD true
+  | xC ⇒ pairT exadecim bool xE false | xD ⇒ pairT exadecim bool xE true
+  | xE ⇒ pairT exadecim bool xF false | xF ⇒ pairT exadecim bool xF true ]
+ | false ⇒ match e with 
+  [ x0 ⇒ pairT exadecim bool x0 false | x1 ⇒ pairT exadecim bool x0 true
+  | x2 ⇒ pairT exadecim bool x1 false | x3 ⇒ pairT exadecim bool x1 true
+  | x4 ⇒ pairT exadecim bool x2 false | x5 ⇒ pairT exadecim bool x2 true
+  | x6 ⇒ pairT exadecim bool x3 false | x7 ⇒ pairT exadecim bool x3 true
+  | x8 ⇒ pairT exadecim bool x4 false | x9 ⇒ pairT exadecim bool x4 true
+  | xA ⇒ pairT exadecim bool x5 false | xB ⇒ pairT exadecim bool x5 true
+  | xC ⇒ pairT exadecim bool x6 false | xD ⇒ pairT exadecim bool x6 true
+  | xE ⇒ pairT exadecim bool x7 false | xF ⇒ pairT exadecim bool x7 true ]
+ ].
+
+(* operatore shift destro *)
+definition shr_ex ≝
+λe:exadecim.match e with 
+ [ x0 ⇒ pairT exadecim bool x0 false | x1 ⇒ pairT exadecim bool x0 true
+ | x2 ⇒ pairT exadecim bool x1 false | x3 ⇒ pairT exadecim bool x1 true
+ | x4 ⇒ pairT exadecim bool x2 false | x5 ⇒ pairT exadecim bool x2 true
+ | x6 ⇒ pairT exadecim bool x3 false | x7 ⇒ pairT exadecim bool x3 true
+ | x8 ⇒ pairT exadecim bool x4 false | x9 ⇒ pairT exadecim bool x4 true
+ | xA ⇒ pairT exadecim bool x5 false | xB ⇒ pairT exadecim bool x5 true
+ | xC ⇒ pairT exadecim bool x6 false | xD ⇒ pairT exadecim bool x6 true
+ | xE ⇒ pairT exadecim bool x7 false | xF ⇒ pairT exadecim bool x7 true ].
+
+(* operatore rotazione destra *)
+definition ror_ex ≝
+λe:exadecim.match e with 
+ [ x0 ⇒ x0 | x1 ⇒ x8 | x2 ⇒ x1 | x3 ⇒ x9
+ | x4 ⇒ x2 | x5 ⇒ xA | x6 ⇒ x3 | x7 ⇒ xB 
+ | x8 ⇒ x4 | x9 ⇒ xC | xA ⇒ x5 | xB ⇒ xD 
+ | xC ⇒ x6 | xD ⇒ xE | xE ⇒ x7 | xF ⇒ xF ].
+
+(* operatore rotazione destra n-volte *)
+let rec ror_ex_n (e:exadecim) (n:nat) on n ≝
+ match n with
+  [ O ⇒ e
+  | S n' ⇒ ror_ex_n (ror_ex e) n' ].
+
+(* operatore rotazione sinistra con carry *)
+definition rcl_ex ≝
+λe:exadecim.λc:bool.match c with
+ [ true ⇒ match e with
+  [ x0 ⇒ pairT exadecim bool x1 false | x1 ⇒ pairT exadecim bool x3 false
+  | x2 ⇒ pairT exadecim bool x5 false | x3 ⇒ pairT exadecim bool x7 false
+  | x4 ⇒ pairT exadecim bool x9 false | x5 ⇒ pairT exadecim bool xB false
+  | x6 ⇒ pairT exadecim bool xD false | x7 ⇒ pairT exadecim bool xF false
+  | x8 ⇒ pairT exadecim bool x1 true  | x9 ⇒ pairT exadecim bool x3 true
+  | xA ⇒ pairT exadecim bool x5 true  | xB ⇒ pairT exadecim bool x7 true
+  | xC ⇒ pairT exadecim bool x9 true  | xD ⇒ pairT exadecim bool xB true
+  | xE ⇒ pairT exadecim bool xD true  | xF ⇒ pairT exadecim bool xF true ]
+ | false ⇒ match e with
+  [ x0 ⇒ pairT exadecim bool x0 false | x1 ⇒ pairT exadecim bool x2 false
+  | x2 ⇒ pairT exadecim bool x4 false | x3 ⇒ pairT exadecim bool x6 false
+  | x4 ⇒ pairT exadecim bool x8 false | x5 ⇒ pairT exadecim bool xA false
+  | x6 ⇒ pairT exadecim bool xC false | x7 ⇒ pairT exadecim bool xE false
+  | x8 ⇒ pairT exadecim bool x0 true  | x9 ⇒ pairT exadecim bool x2 true
+  | xA ⇒ pairT exadecim bool x4 true  | xB ⇒ pairT exadecim bool x6 true
+  | xC ⇒ pairT exadecim bool x8 true  | xD ⇒ pairT exadecim bool xA true
+  | xE ⇒ pairT exadecim bool xC true  | xF ⇒ pairT exadecim bool xE true ]
+ ].
+
+(* operatore shift sinistro *)
+definition shl_ex ≝
+λe:exadecim.match e with
+ [ x0 ⇒ pairT exadecim bool x0 false | x1 ⇒ pairT exadecim bool x2 false
+ | x2 ⇒ pairT exadecim bool x4 false | x3 ⇒ pairT exadecim bool x6 false
+ | x4 ⇒ pairT exadecim bool x8 false | x5 ⇒ pairT exadecim bool xA false
+ | x6 ⇒ pairT exadecim bool xC false | x7 ⇒ pairT exadecim bool xE false
+ | x8 ⇒ pairT exadecim bool x0 true  | x9 ⇒ pairT exadecim bool x2 true
+ | xA ⇒ pairT exadecim bool x4 true  | xB ⇒ pairT exadecim bool x6 true
+ | xC ⇒ pairT exadecim bool x8 true  | xD ⇒ pairT exadecim bool xA true
+ | xE ⇒ pairT exadecim bool xC true  | xF ⇒ pairT exadecim bool xE true ].
+
+(* operatore rotazione sinistra *)
+definition rol_ex ≝
+λe:exadecim.match e with
+ [ x0 ⇒ x0 | x1 ⇒ x2 | x2 ⇒ x4 | x3 ⇒ x6
+ | x4 ⇒ x8 | x5 ⇒ xA | x6 ⇒ xC | x7 ⇒ xE
+ | x8 ⇒ x1 | x9 ⇒ x3 | xA ⇒ x5 | xB ⇒ x7
+ | xC ⇒ x9 | xD ⇒ xB | xE ⇒ xD | xF ⇒ xF ].
+
+(* operatore rotazione sinistra n-volte *)
+let rec rol_ex_n (e:exadecim) (n:nat) on n ≝
+ match n with
+  [ O ⇒ e
+  | S n' ⇒ rol_ex_n (rol_ex e) n' ].
+
+(* operatore not/complemento a 1 *)
+definition not_ex ≝
+λe:exadecim.match e with
+ [ x0 ⇒ xF | x1 ⇒ xE | x2 ⇒ xD | x3 ⇒ xC
+ | x4 ⇒ xB | x5 ⇒ xA | x6 ⇒ x9 | x7 ⇒ x8
+ | x8 ⇒ x7 | x9 ⇒ x6 | xA ⇒ x5 | xB ⇒ x4
+ | xC ⇒ x3 | xD ⇒ x2 | xE ⇒ x1 | xF ⇒ x0 ].
+
+(* operatore somma con carry *)
+definition plus_ex ≝
+λe1,e2:exadecim.λc:bool.
+  match c with
+   [ true ⇒
+      match e1 with
+       [ x0 ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool x1 false
+            | x1 ⇒ pairT exadecim bool x2 false
+            | x2 ⇒ pairT exadecim bool x3 false
+            | x3 ⇒ pairT exadecim bool x4 false
+            | x4 ⇒ pairT exadecim bool x5 false
+            | x5 ⇒ pairT exadecim bool x6 false
+            | x6 ⇒ pairT exadecim bool x7 false
+            | x7 ⇒ pairT exadecim bool x8 false
+            | x8 ⇒ pairT exadecim bool x9 false
+            | x9 ⇒ pairT exadecim bool xA false
+            | xA ⇒ pairT exadecim bool xB false
+            | xB ⇒ pairT exadecim bool xC false
+            | xC ⇒ pairT exadecim bool xD false
+            | xD ⇒ pairT exadecim bool xE false
+            | xE ⇒ pairT exadecim bool xF false
+            | xF ⇒ pairT exadecim bool x0 true ] 
+       | x1 ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool x2 false
+            | x1 ⇒ pairT exadecim bool x3 false
+            | x2 ⇒ pairT exadecim bool x4 false
+            | x3 ⇒ pairT exadecim bool x5 false
+            | x4 ⇒ pairT exadecim bool x6 false
+            | x5 ⇒ pairT exadecim bool x7 false
+            | x6 ⇒ pairT exadecim bool x8 false
+            | x7 ⇒ pairT exadecim bool x9 false
+            | x8 ⇒ pairT exadecim bool xA false
+            | x9 ⇒ pairT exadecim bool xB false
+            | xA ⇒ pairT exadecim bool xC false
+            | xB ⇒ pairT exadecim bool xD false
+            | xC ⇒ pairT exadecim bool xE false
+            | xD ⇒ pairT exadecim bool xF false
+            | xE ⇒ pairT exadecim bool x0 true
+            | xF ⇒ pairT exadecim bool x1 true ] 
+       | x2 ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool x3 false
+            | x1 ⇒ pairT exadecim bool x4 false
+            | x2 ⇒ pairT exadecim bool x5 false
+            | x3 ⇒ pairT exadecim bool x6 false
+            | x4 ⇒ pairT exadecim bool x7 false
+            | x5 ⇒ pairT exadecim bool x8 false
+            | x6 ⇒ pairT exadecim bool x9 false
+            | x7 ⇒ pairT exadecim bool xA false
+            | x8 ⇒ pairT exadecim bool xB false
+            | x9 ⇒ pairT exadecim bool xC false
+            | xA ⇒ pairT exadecim bool xD false
+            | xB ⇒ pairT exadecim bool xE false
+            | xC ⇒ pairT exadecim bool xF false
+            | xD ⇒ pairT exadecim bool x0 true
+            | xE ⇒ pairT exadecim bool x1 true
+            | xF ⇒ pairT exadecim bool x2 true ] 
+       | x3 ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool x4 false
+            | x1 ⇒ pairT exadecim bool x5 false
+            | x2 ⇒ pairT exadecim bool x6 false
+            | x3 ⇒ pairT exadecim bool x7 false
+            | x4 ⇒ pairT exadecim bool x8 false
+            | x5 ⇒ pairT exadecim bool x9 false
+            | x6 ⇒ pairT exadecim bool xA false
+            | x7 ⇒ pairT exadecim bool xB false
+            | x8 ⇒ pairT exadecim bool xC false
+            | x9 ⇒ pairT exadecim bool xD false
+            | xA ⇒ pairT exadecim bool xE false
+            | xB ⇒ pairT exadecim bool xF false
+            | xC ⇒ pairT exadecim bool x0 true
+            | xD ⇒ pairT exadecim bool x1 true
+            | xE ⇒ pairT exadecim bool x2 true
+            | xF ⇒ pairT exadecim bool x3 true ] 
+       | x4 ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool x5 false
+            | x1 ⇒ pairT exadecim bool x6 false
+            | x2 ⇒ pairT exadecim bool x7 false
+            | x3 ⇒ pairT exadecim bool x8 false
+            | x4 ⇒ pairT exadecim bool x9 false
+            | x5 ⇒ pairT exadecim bool xA false
+            | x6 ⇒ pairT exadecim bool xB false
+            | x7 ⇒ pairT exadecim bool xC false
+            | x8 ⇒ pairT exadecim bool xD false
+            | x9 ⇒ pairT exadecim bool xE false
+            | xA ⇒ pairT exadecim bool xF false
+            | xB ⇒ pairT exadecim bool x0 true
+            | xC ⇒ pairT exadecim bool x1 true
+            | xD ⇒ pairT exadecim bool x2 true
+            | xE ⇒ pairT exadecim bool x3 true
+            | xF ⇒ pairT exadecim bool x4 true ] 
+       | x5 ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool x6 false
+            | x1 ⇒ pairT exadecim bool x7 false
+            | x2 ⇒ pairT exadecim bool x8 false
+            | x3 ⇒ pairT exadecim bool x9 false
+            | x4 ⇒ pairT exadecim bool xA false
+            | x5 ⇒ pairT exadecim bool xB false
+            | x6 ⇒ pairT exadecim bool xC false
+            | x7 ⇒ pairT exadecim bool xD false
+            | x8 ⇒ pairT exadecim bool xE false
+            | x9 ⇒ pairT exadecim bool xF false
+            | xA ⇒ pairT exadecim bool x0 true
+            | xB ⇒ pairT exadecim bool x1 true
+            | xC ⇒ pairT exadecim bool x2 true
+            | xD ⇒ pairT exadecim bool x3 true
+            | xE ⇒ pairT exadecim bool x4 true
+            | xF ⇒ pairT exadecim bool x5 true ] 
+       | x6 ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool x7 false
+            | x1 ⇒ pairT exadecim bool x8 false
+            | x2 ⇒ pairT exadecim bool x9 false
+            | x3 ⇒ pairT exadecim bool xA false
+            | x4 ⇒ pairT exadecim bool xB false
+            | x5 ⇒ pairT exadecim bool xC false
+            | x6 ⇒ pairT exadecim bool xD false
+            | x7 ⇒ pairT exadecim bool xE false
+            | x8 ⇒ pairT exadecim bool xF false
+            | x9 ⇒ pairT exadecim bool x0 true
+            | xA ⇒ pairT exadecim bool x1 true
+            | xB ⇒ pairT exadecim bool x2 true
+            | xC ⇒ pairT exadecim bool x3 true
+            | xD ⇒ pairT exadecim bool x4 true
+            | xE ⇒ pairT exadecim bool x5 true
+            | xF ⇒ pairT exadecim bool x6 true ] 
+       | x7 ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool x8 false
+            | x1 ⇒ pairT exadecim bool x9 false
+            | x2 ⇒ pairT exadecim bool xA false
+            | x3 ⇒ pairT exadecim bool xB false
+            | x4 ⇒ pairT exadecim bool xC false
+            | x5 ⇒ pairT exadecim bool xD false
+            | x6 ⇒ pairT exadecim bool xE false
+            | x7 ⇒ pairT exadecim bool xF false
+            | x8 ⇒ pairT exadecim bool x0 true
+            | x9 ⇒ pairT exadecim bool x1 true
+            | xA ⇒ pairT exadecim bool x2 true
+            | xB ⇒ pairT exadecim bool x3 true
+            | xC ⇒ pairT exadecim bool x4 true
+            | xD ⇒ pairT exadecim bool x5 true
+            | xE ⇒ pairT exadecim bool x6 true
+            | xF ⇒ pairT exadecim bool x7 true ] 
+       | x8 ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool x9 false
+            | x1 ⇒ pairT exadecim bool xA false
+            | x2 ⇒ pairT exadecim bool xB false
+            | x3 ⇒ pairT exadecim bool xC false
+            | x4 ⇒ pairT exadecim bool xD false
+            | x5 ⇒ pairT exadecim bool xE false
+            | x6 ⇒ pairT exadecim bool xF false
+            | x7 ⇒ pairT exadecim bool x0 true
+            | x8 ⇒ pairT exadecim bool x1 true
+            | x9 ⇒ pairT exadecim bool x2 true
+            | xA ⇒ pairT exadecim bool x3 true
+            | xB ⇒ pairT exadecim bool x4 true
+            | xC ⇒ pairT exadecim bool x5 true
+            | xD ⇒ pairT exadecim bool x6 true
+            | xE ⇒ pairT exadecim bool x7 true
+            | xF ⇒ pairT exadecim bool x8 true ] 
+       | x9 ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool xA false
+            | x1 ⇒ pairT exadecim bool xB false
+            | x2 ⇒ pairT exadecim bool xC false
+            | x3 ⇒ pairT exadecim bool xD false
+            | x4 ⇒ pairT exadecim bool xE false
+            | x5 ⇒ pairT exadecim bool xF false
+            | x6 ⇒ pairT exadecim bool x0 true
+            | x7 ⇒ pairT exadecim bool x1 true
+            | x8 ⇒ pairT exadecim bool x2 true
+            | x9 ⇒ pairT exadecim bool x3 true
+            | xA ⇒ pairT exadecim bool x4 true
+            | xB ⇒ pairT exadecim bool x5 true
+            | xC ⇒ pairT exadecim bool x6 true
+            | xD ⇒ pairT exadecim bool x7 true
+            | xE ⇒ pairT exadecim bool x8 true
+            | xF ⇒ pairT exadecim bool x9 true ] 
+       | xA ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool xB false
+            | x1 ⇒ pairT exadecim bool xC false
+            | x2 ⇒ pairT exadecim bool xD false
+            | x3 ⇒ pairT exadecim bool xE false
+            | x4 ⇒ pairT exadecim bool xF false
+            | x5 ⇒ pairT exadecim bool x0 true
+            | x6 ⇒ pairT exadecim bool x1 true
+            | x7 ⇒ pairT exadecim bool x2 true
+            | x8 ⇒ pairT exadecim bool x3 true
+            | x9 ⇒ pairT exadecim bool x4 true
+            | xA ⇒ pairT exadecim bool x5 true
+            | xB ⇒ pairT exadecim bool x6 true
+            | xC ⇒ pairT exadecim bool x7 true
+            | xD ⇒ pairT exadecim bool x8 true
+            | xE ⇒ pairT exadecim bool x9 true
+            | xF ⇒ pairT exadecim bool xA true ] 
+       | xB ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool xC false
+            | x1 ⇒ pairT exadecim bool xD false
+            | x2 ⇒ pairT exadecim bool xE false
+            | x3 ⇒ pairT exadecim bool xF false
+            | x4 ⇒ pairT exadecim bool x0 true
+            | x5 ⇒ pairT exadecim bool x1 true
+            | x6 ⇒ pairT exadecim bool x2 true
+            | x7 ⇒ pairT exadecim bool x3 true
+            | x8 ⇒ pairT exadecim bool x4 true
+            | x9 ⇒ pairT exadecim bool x5 true
+            | xA ⇒ pairT exadecim bool x6 true
+            | xB ⇒ pairT exadecim bool x7 true
+            | xC ⇒ pairT exadecim bool x8 true
+            | xD ⇒ pairT exadecim bool x9 true
+            | xE ⇒ pairT exadecim bool xA true
+            | xF ⇒ pairT exadecim bool xB true ] 
+       | xC ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool xD false
+            | x1 ⇒ pairT exadecim bool xE false
+            | x2 ⇒ pairT exadecim bool xF false
+            | x3 ⇒ pairT exadecim bool x0 true
+            | x4 ⇒ pairT exadecim bool x1 true
+            | x5 ⇒ pairT exadecim bool x2 true
+            | x6 ⇒ pairT exadecim bool x3 true
+            | x7 ⇒ pairT exadecim bool x4 true
+            | x8 ⇒ pairT exadecim bool x5 true
+            | x9 ⇒ pairT exadecim bool x6 true
+            | xA ⇒ pairT exadecim bool x7 true
+            | xB ⇒ pairT exadecim bool x8 true
+            | xC ⇒ pairT exadecim bool x9 true
+            | xD ⇒ pairT exadecim bool xA true
+            | xE ⇒ pairT exadecim bool xB true
+            | xF ⇒ pairT exadecim bool xC true ] 
+       | xD ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool xE false
+            | x1 ⇒ pairT exadecim bool xF false
+            | x2 ⇒ pairT exadecim bool x0 true
+            | x3 ⇒ pairT exadecim bool x1 true
+            | x4 ⇒ pairT exadecim bool x2 true
+            | x5 ⇒ pairT exadecim bool x3 true
+            | x6 ⇒ pairT exadecim bool x4 true
+            | x7 ⇒ pairT exadecim bool x5 true
+            | x8 ⇒ pairT exadecim bool x6 true
+            | x9 ⇒ pairT exadecim bool x7 true
+            | xA ⇒ pairT exadecim bool x8 true
+            | xB ⇒ pairT exadecim bool x9 true
+            | xC ⇒ pairT exadecim bool xA true
+            | xD ⇒ pairT exadecim bool xB true
+            | xE ⇒ pairT exadecim bool xC true
+            | xF ⇒ pairT exadecim bool xD true ] 
+       | xE ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool xF false
+            | x1 ⇒ pairT exadecim bool x0 true
+            | x2 ⇒ pairT exadecim bool x1 true
+            | x3 ⇒ pairT exadecim bool x2 true
+            | x4 ⇒ pairT exadecim bool x3 true
+            | x5 ⇒ pairT exadecim bool x4 true
+            | x6 ⇒ pairT exadecim bool x5 true
+            | x7 ⇒ pairT exadecim bool x6 true
+            | x8 ⇒ pairT exadecim bool x7 true
+            | x9 ⇒ pairT exadecim bool x8 true
+            | xA ⇒ pairT exadecim bool x9 true
+            | xB ⇒ pairT exadecim bool xA true
+            | xC ⇒ pairT exadecim bool xB true
+            | xD ⇒ pairT exadecim bool xC true
+            | xE ⇒ pairT exadecim bool xD true
+            | xF ⇒ pairT exadecim bool xE true ]
+       | xF ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool x0 true
+            | x1 ⇒ pairT exadecim bool x1 true
+            | x2 ⇒ pairT exadecim bool x2 true
+            | x3 ⇒ pairT exadecim bool x3 true
+            | x4 ⇒ pairT exadecim bool x4 true
+            | x5 ⇒ pairT exadecim bool x5 true
+            | x6 ⇒ pairT exadecim bool x6 true
+            | x7 ⇒ pairT exadecim bool x7 true
+            | x8 ⇒ pairT exadecim bool x8 true
+            | x9 ⇒ pairT exadecim bool x9 true
+            | xA ⇒ pairT exadecim bool xA true
+            | xB ⇒ pairT exadecim bool xB true
+            | xC ⇒ pairT exadecim bool xC true
+            | xD ⇒ pairT exadecim bool xD true
+            | xE ⇒ pairT exadecim bool xE true
+            | xF ⇒ pairT exadecim bool xF true ] 
+       ]
+   | false ⇒
+      match e1 with
+       [ x0 ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool x0 false
+            | x1 ⇒ pairT exadecim bool x1 false
+            | x2 ⇒ pairT exadecim bool x2 false
+            | x3 ⇒ pairT exadecim bool x3 false
+            | x4 ⇒ pairT exadecim bool x4 false
+            | x5 ⇒ pairT exadecim bool x5 false
+            | x6 ⇒ pairT exadecim bool x6 false
+            | x7 ⇒ pairT exadecim bool x7 false
+            | x8 ⇒ pairT exadecim bool x8 false
+            | x9 ⇒ pairT exadecim bool x9 false
+            | xA ⇒ pairT exadecim bool xA false
+            | xB ⇒ pairT exadecim bool xB false
+            | xC ⇒ pairT exadecim bool xC false
+            | xD ⇒ pairT exadecim bool xD false
+            | xE ⇒ pairT exadecim bool xE false
+            | xF ⇒ pairT exadecim bool xF false ] 
+       | x1 ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool x1 false
+            | x1 ⇒ pairT exadecim bool x2 false
+            | x2 ⇒ pairT exadecim bool x3 false
+            | x3 ⇒ pairT exadecim bool x4 false
+            | x4 ⇒ pairT exadecim bool x5 false
+            | x5 ⇒ pairT exadecim bool x6 false
+            | x6 ⇒ pairT exadecim bool x7 false
+            | x7 ⇒ pairT exadecim bool x8 false
+            | x8 ⇒ pairT exadecim bool x9 false
+            | x9 ⇒ pairT exadecim bool xA false
+            | xA ⇒ pairT exadecim bool xB false
+            | xB ⇒ pairT exadecim bool xC false
+            | xC ⇒ pairT exadecim bool xD false
+            | xD ⇒ pairT exadecim bool xE false
+            | xE ⇒ pairT exadecim bool xF false
+            | xF ⇒ pairT exadecim bool x0 true ] 
+       | x2 ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool x2 false
+            | x1 ⇒ pairT exadecim bool x3 false
+            | x2 ⇒ pairT exadecim bool x4 false
+            | x3 ⇒ pairT exadecim bool x5 false
+            | x4 ⇒ pairT exadecim bool x6 false
+            | x5 ⇒ pairT exadecim bool x7 false
+            | x6 ⇒ pairT exadecim bool x8 false
+            | x7 ⇒ pairT exadecim bool x9 false
+            | x8 ⇒ pairT exadecim bool xA false
+            | x9 ⇒ pairT exadecim bool xB false
+            | xA ⇒ pairT exadecim bool xC false
+            | xB ⇒ pairT exadecim bool xD false
+            | xC ⇒ pairT exadecim bool xE false
+            | xD ⇒ pairT exadecim bool xF false
+            | xE ⇒ pairT exadecim bool x0 true
+            | xF ⇒ pairT exadecim bool x1 true ] 
+       | x3 ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool x3 false
+            | x1 ⇒ pairT exadecim bool x4 false
+            | x2 ⇒ pairT exadecim bool x5 false
+            | x3 ⇒ pairT exadecim bool x6 false
+            | x4 ⇒ pairT exadecim bool x7 false
+            | x5 ⇒ pairT exadecim bool x8 false
+            | x6 ⇒ pairT exadecim bool x9 false
+            | x7 ⇒ pairT exadecim bool xA false
+            | x8 ⇒ pairT exadecim bool xB false
+            | x9 ⇒ pairT exadecim bool xC false
+            | xA ⇒ pairT exadecim bool xD false
+            | xB ⇒ pairT exadecim bool xE false
+            | xC ⇒ pairT exadecim bool xF false
+            | xD ⇒ pairT exadecim bool x0 true
+            | xE ⇒ pairT exadecim bool x1 true
+            | xF ⇒ pairT exadecim bool x2 true ] 
+       | x4 ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool x4 false
+            | x1 ⇒ pairT exadecim bool x5 false
+            | x2 ⇒ pairT exadecim bool x6 false
+            | x3 ⇒ pairT exadecim bool x7 false
+            | x4 ⇒ pairT exadecim bool x8 false
+            | x5 ⇒ pairT exadecim bool x9 false
+            | x6 ⇒ pairT exadecim bool xA false
+            | x7 ⇒ pairT exadecim bool xB false
+            | x8 ⇒ pairT exadecim bool xC false
+            | x9 ⇒ pairT exadecim bool xD false
+            | xA ⇒ pairT exadecim bool xE false
+            | xB ⇒ pairT exadecim bool xF false
+            | xC ⇒ pairT exadecim bool x0 true
+            | xD ⇒ pairT exadecim bool x1 true
+            | xE ⇒ pairT exadecim bool x2 true
+            | xF ⇒ pairT exadecim bool x3 true ] 
+       | x5 ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool x5 false
+            | x1 ⇒ pairT exadecim bool x6 false
+            | x2 ⇒ pairT exadecim bool x7 false
+            | x3 ⇒ pairT exadecim bool x8 false
+            | x4 ⇒ pairT exadecim bool x9 false
+            | x5 ⇒ pairT exadecim bool xA false
+            | x6 ⇒ pairT exadecim bool xB false
+            | x7 ⇒ pairT exadecim bool xC false
+            | x8 ⇒ pairT exadecim bool xD false
+            | x9 ⇒ pairT exadecim bool xE false
+            | xA ⇒ pairT exadecim bool xF false
+            | xB ⇒ pairT exadecim bool x0 true
+            | xC ⇒ pairT exadecim bool x1 true
+            | xD ⇒ pairT exadecim bool x2 true
+            | xE ⇒ pairT exadecim bool x3 true
+            | xF ⇒ pairT exadecim bool x4 true ] 
+       | x6 ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool x6 false
+            | x1 ⇒ pairT exadecim bool x7 false
+            | x2 ⇒ pairT exadecim bool x8 false
+            | x3 ⇒ pairT exadecim bool x9 false
+            | x4 ⇒ pairT exadecim bool xA false
+            | x5 ⇒ pairT exadecim bool xB false
+            | x6 ⇒ pairT exadecim bool xC false
+            | x7 ⇒ pairT exadecim bool xD false
+            | x8 ⇒ pairT exadecim bool xE false
+            | x9 ⇒ pairT exadecim bool xF false
+            | xA ⇒ pairT exadecim bool x0 true
+            | xB ⇒ pairT exadecim bool x1 true
+            | xC ⇒ pairT exadecim bool x2 true
+            | xD ⇒ pairT exadecim bool x3 true
+            | xE ⇒ pairT exadecim bool x4 true
+            | xF ⇒ pairT exadecim bool x5 true ] 
+       | x7 ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool x7 false
+            | x1 ⇒ pairT exadecim bool x8 false
+            | x2 ⇒ pairT exadecim bool x9 false
+            | x3 ⇒ pairT exadecim bool xA false
+            | x4 ⇒ pairT exadecim bool xB false
+            | x5 ⇒ pairT exadecim bool xC false
+            | x6 ⇒ pairT exadecim bool xD false
+            | x7 ⇒ pairT exadecim bool xE false
+            | x8 ⇒ pairT exadecim bool xF false
+            | x9 ⇒ pairT exadecim bool x0 true
+            | xA ⇒ pairT exadecim bool x1 true
+            | xB ⇒ pairT exadecim bool x2 true
+            | xC ⇒ pairT exadecim bool x3 true
+            | xD ⇒ pairT exadecim bool x4 true
+            | xE ⇒ pairT exadecim bool x5 true
+            | xF ⇒ pairT exadecim bool x6 true ] 
+       | x8 ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool x8 false
+            | x1 ⇒ pairT exadecim bool x9 false
+            | x2 ⇒ pairT exadecim bool xA false
+            | x3 ⇒ pairT exadecim bool xB false
+            | x4 ⇒ pairT exadecim bool xC false
+            | x5 ⇒ pairT exadecim bool xD false
+            | x6 ⇒ pairT exadecim bool xE false
+            | x7 ⇒ pairT exadecim bool xF false
+            | x8 ⇒ pairT exadecim bool x0 true
+            | x9 ⇒ pairT exadecim bool x1 true
+            | xA ⇒ pairT exadecim bool x2 true
+            | xB ⇒ pairT exadecim bool x3 true
+            | xC ⇒ pairT exadecim bool x4 true
+            | xD ⇒ pairT exadecim bool x5 true
+            | xE ⇒ pairT exadecim bool x6 true
+            | xF ⇒ pairT exadecim bool x7 true ] 
+       | x9 ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool x9 false
+            | x1 ⇒ pairT exadecim bool xA false
+            | x2 ⇒ pairT exadecim bool xB false
+            | x3 ⇒ pairT exadecim bool xC false
+            | x4 ⇒ pairT exadecim bool xD false
+            | x5 ⇒ pairT exadecim bool xE false
+            | x6 ⇒ pairT exadecim bool xF false
+            | x7 ⇒ pairT exadecim bool x0 true
+            | x8 ⇒ pairT exadecim bool x1 true
+            | x9 ⇒ pairT exadecim bool x2 true
+            | xA ⇒ pairT exadecim bool x3 true
+            | xB ⇒ pairT exadecim bool x4 true
+            | xC ⇒ pairT exadecim bool x5 true
+            | xD ⇒ pairT exadecim bool x6 true
+            | xE ⇒ pairT exadecim bool x7 true
+            | xF ⇒ pairT exadecim bool x8 true ] 
+       | xA ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool xA false
+            | x1 ⇒ pairT exadecim bool xB false
+            | x2 ⇒ pairT exadecim bool xC false
+            | x3 ⇒ pairT exadecim bool xD false
+            | x4 ⇒ pairT exadecim bool xE false
+            | x5 ⇒ pairT exadecim bool xF false
+            | x6 ⇒ pairT exadecim bool x0 true
+            | x7 ⇒ pairT exadecim bool x1 true
+            | x8 ⇒ pairT exadecim bool x2 true
+            | x9 ⇒ pairT exadecim bool x3 true
+            | xA ⇒ pairT exadecim bool x4 true
+            | xB ⇒ pairT exadecim bool x5 true
+            | xC ⇒ pairT exadecim bool x6 true
+            | xD ⇒ pairT exadecim bool x7 true
+            | xE ⇒ pairT exadecim bool x8 true
+            | xF ⇒ pairT exadecim bool x9 true ] 
+       | xB ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool xB false
+            | x1 ⇒ pairT exadecim bool xC false
+            | x2 ⇒ pairT exadecim bool xD false
+            | x3 ⇒ pairT exadecim bool xE false
+            | x4 ⇒ pairT exadecim bool xF false
+            | x5 ⇒ pairT exadecim bool x0 true
+            | x6 ⇒ pairT exadecim bool x1 true
+            | x7 ⇒ pairT exadecim bool x2 true
+            | x8 ⇒ pairT exadecim bool x3 true
+            | x9 ⇒ pairT exadecim bool x4 true
+            | xA ⇒ pairT exadecim bool x5 true
+            | xB ⇒ pairT exadecim bool x6 true
+            | xC ⇒ pairT exadecim bool x7 true
+            | xD ⇒ pairT exadecim bool x8 true
+            | xE ⇒ pairT exadecim bool x9 true
+            | xF ⇒ pairT exadecim bool xA true ] 
+       | xC ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool xC false
+            | x1 ⇒ pairT exadecim bool xD false
+            | x2 ⇒ pairT exadecim bool xE false
+            | x3 ⇒ pairT exadecim bool xF false
+            | x4 ⇒ pairT exadecim bool x0 true
+            | x5 ⇒ pairT exadecim bool x1 true
+            | x6 ⇒ pairT exadecim bool x2 true
+            | x7 ⇒ pairT exadecim bool x3 true
+            | x8 ⇒ pairT exadecim bool x4 true
+            | x9 ⇒ pairT exadecim bool x5 true
+            | xA ⇒ pairT exadecim bool x6 true
+            | xB ⇒ pairT exadecim bool x7 true
+            | xC ⇒ pairT exadecim bool x8 true
+            | xD ⇒ pairT exadecim bool x9 true
+            | xE ⇒ pairT exadecim bool xA true
+            | xF ⇒ pairT exadecim bool xB true ] 
+       | xD ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool xD false
+            | x1 ⇒ pairT exadecim bool xE false
+            | x2 ⇒ pairT exadecim bool xF false
+            | x3 ⇒ pairT exadecim bool x0 true
+            | x4 ⇒ pairT exadecim bool x1 true
+            | x5 ⇒ pairT exadecim bool x2 true
+            | x6 ⇒ pairT exadecim bool x3 true
+            | x7 ⇒ pairT exadecim bool x4 true
+            | x8 ⇒ pairT exadecim bool x5 true
+            | x9 ⇒ pairT exadecim bool x6 true
+            | xA ⇒ pairT exadecim bool x7 true
+            | xB ⇒ pairT exadecim bool x8 true
+            | xC ⇒ pairT exadecim bool x9 true
+            | xD ⇒ pairT exadecim bool xA true
+            | xE ⇒ pairT exadecim bool xB true
+            | xF ⇒ pairT exadecim bool xC true ] 
+       | xE ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool xE false
+            | x1 ⇒ pairT exadecim bool xF false
+            | x2 ⇒ pairT exadecim bool x0 true
+            | x3 ⇒ pairT exadecim bool x1 true
+            | x4 ⇒ pairT exadecim bool x2 true
+            | x5 ⇒ pairT exadecim bool x3 true
+            | x6 ⇒ pairT exadecim bool x4 true
+            | x7 ⇒ pairT exadecim bool x5 true
+            | x8 ⇒ pairT exadecim bool x6 true
+            | x9 ⇒ pairT exadecim bool x7 true
+            | xA ⇒ pairT exadecim bool x8 true
+            | xB ⇒ pairT exadecim bool x9 true
+            | xC ⇒ pairT exadecim bool xA true
+            | xD ⇒ pairT exadecim bool xB true
+            | xE ⇒ pairT exadecim bool xC true
+            | xF ⇒ pairT exadecim bool xD true ] 
+       | xF ⇒
+           match e2 with
+            [ x0 ⇒ pairT exadecim bool xF false
+            | x1 ⇒ pairT exadecim bool x0 true
+            | x2 ⇒ pairT exadecim bool x1 true
+            | x3 ⇒ pairT exadecim bool x2 true
+            | x4 ⇒ pairT exadecim bool x3 true
+            | x5 ⇒ pairT exadecim bool x4 true
+            | x6 ⇒ pairT exadecim bool x5 true
+            | x7 ⇒ pairT exadecim bool x6 true
+            | x8 ⇒ pairT exadecim bool x7 true
+            | x9 ⇒ pairT exadecim bool x8 true
+            | xA ⇒ pairT exadecim bool x9 true
+            | xB ⇒ pairT exadecim bool xA true
+            | xC ⇒ pairT exadecim bool xB true
+            | xD ⇒ pairT exadecim bool xC true
+            | xE ⇒ pairT exadecim bool xD true
+            | xF ⇒ pairT exadecim bool xE true ]
+       ]
+   ].
+
+(* operatore somma senza carry *)
+definition plus_exnc ≝
+λe1,e2:exadecim.fstT ?? (plus_ex e1 e2 false).
+
+(* operatore carry della somma *)
+definition plus_exc ≝
+λe1,e2:exadecim.sndT ?? (plus_ex e1 e2 false).
+
+(* operatore Most Significant Bit *)
+definition MSB_ex ≝
+λe:exadecim.match e with
+ [ x0 ⇒ false | x1 ⇒ false | x2 ⇒ false | x3 ⇒ false
+ | x4 ⇒ false | x5 ⇒ false | x6 ⇒ false | x7 ⇒ false
+ | x8 ⇒ true  | x9 ⇒ true  | xA ⇒ true  | xB ⇒ true
+ | xC ⇒ true  | xD ⇒ true  | xE ⇒ true  | xF ⇒ true ].
+
+(* esadecimali → naturali *)
+definition nat_of_exadecim ≝
+λe:exadecim.
+ match e with
+  [ x0 ⇒ 0 | x1 ⇒ 1 | x2 ⇒ 2  | x3 ⇒ 3  | x4 ⇒ 4  | x5 ⇒ 5  | x6 ⇒ 6  | x7 ⇒ 7
+  | x8 ⇒ 8 | x9 ⇒ 9 | xA ⇒ 10 | xB ⇒ 11 | xC ⇒ 12 | xD ⇒ 13 | xE ⇒ 14 | xF ⇒ 15 ].
+
+coercion cic:/matita/freescale/exadecim/nat_of_exadecim.con.
+
+(* naturali → esadecimali *)
+let rec exadecim_of_nat n ≝
+ match n with [ O ⇒ x0 | S n ⇒
+  match n with [ O ⇒ x1 | S n ⇒
+   match n with [ O ⇒ x2 | S n ⇒ 
+    match n with [ O ⇒ x3 | S n ⇒ 
+     match n with [ O ⇒ x4 | S n ⇒ 
+      match n with [ O ⇒ x5 | S n ⇒ 
+       match n with [ O ⇒ x6 | S n ⇒ 
+        match n with [ O ⇒ x7 | S n ⇒ 
+         match n with [ O ⇒ x8 | S n ⇒ 
+          match n with [ O ⇒ x9 | S n ⇒ 
+           match n with [ O ⇒ xA | S n ⇒ 
+            match n with [ O ⇒ xB | S n ⇒ 
+             match n with [ O ⇒ xC | S n ⇒ 
+              match n with [ O ⇒ xD | S n ⇒ 
+               match n with [ O ⇒ xE | S n ⇒ 
+                match n with [ O ⇒ xF | S n ⇒ exadecim_of_nat n ]]]]]]]]]]]]]]]]. 
+
+(* operatore predecessore *)
+definition pred_ex ≝
+λe:exadecim.
+ match e with
+  [ x0 ⇒ xF | x1 ⇒ x0 | x2 ⇒ x1 | x3 ⇒ x2 | x4 ⇒ x3 | x5 ⇒ x4 | x6 ⇒ x5 | x7 ⇒ x6
+  | x8 ⇒ x7 | x9 ⇒ x8 | xA ⇒ x9 | xB ⇒ xA | xC ⇒ xB | xD ⇒ xC | xE ⇒ xD | xF ⇒ xE ].
+
+(* operatore successore *)
+definition succ_ex ≝
+λe:exadecim.
+ match e with
+  [ x0 ⇒ x1 | x1 ⇒ x2 | x2 ⇒ x3 | x3 ⇒ x4 | x4 ⇒ x5 | x5 ⇒ x6 | x6 ⇒ x7 | x7 ⇒ x8
+  | x8 ⇒ x9 | x9 ⇒ xA | xA ⇒ xB | xB ⇒ xC | xC ⇒ xD | xD ⇒ xE | xE ⇒ xF | xF ⇒ x0 ].
+
+(* operatore neg/complemento a 2 *)
+definition compl_ex ≝
+λe:exadecim.match e with
+ [ x0 ⇒ x0 | x1 ⇒ xF | x2 ⇒ xE | x3 ⇒ xD
+ | x4 ⇒ xC | x5 ⇒ xB | x6 ⇒ xA | x7 ⇒ x9
+ | x8 ⇒ x8 | x9 ⇒ x7 | xA ⇒ x6 | xB ⇒ x5
+ | xC ⇒ x4 | xD ⇒ x3 | xE ⇒ x2 | xF ⇒ x1 ].
+
+(* iteratore sugli esadecimali *)
+definition forall_exadecim ≝ λP.
+ P x0 ⊗ P x1 ⊗ P x2 ⊗ P x3 ⊗ P x4 ⊗ P x5 ⊗ P x6 ⊗ P x7 ⊗
+ P x8 ⊗ P x9 ⊗ P xA ⊗ P xB ⊗ P xC ⊗ P xD ⊗ P xE ⊗ P xF.
+
+(* ********************** *)
+(* TEOREMI/LEMMMI/ASSIOMI *)
+(* ********************** *)
+
+(* ESPERIMENTO UTILIZZO DELL'ITERATORE *)
+
+(*
+lemma forall_exadecim_eqProp_left_aux :
+ ∀P.(∀e:exadecim.P e = true) →
+ ((P x0 ⊗ P x1 ⊗ P x2 ⊗ P x3 ⊗ P x4 ⊗ P x5 ⊗ P x6 ⊗ P x7 ⊗
+   P x8 ⊗ P x9 ⊗ P xA ⊗ P xB ⊗ P xC ⊗ P xD ⊗ P xE ⊗ P xF) = true).
+ intros;
+ elim H; elim P; [ 2: reflexivity; ] elim H; elim P; [ 2: reflexivity; ]
+ elim H; elim P; [ 2: reflexivity; ] elim H; elim P; [ 2: reflexivity; ]
+ elim H; elim P; [ 2: reflexivity; ] elim H; elim P; [ 2: reflexivity; ]
+ elim H; elim P; [ 2: reflexivity; ] elim H; elim P; [ 2: reflexivity; ]
+ elim H; elim P; [ 2: reflexivity; ] elim H; elim P; [ 2: reflexivity; ]
+ elim H; elim P; [ 2: reflexivity; ] elim H; elim P; [ 2: reflexivity; ]
+ elim H; elim P; [ 2: reflexivity; ] elim H; elim P; [ 2: reflexivity; ]
+ elim H; elim P; [ 2: reflexivity; ] elim H; elim P; [ 2: reflexivity; ]
+ reflexivity.
+qed.
+
+lemma forall_exadecim_eqProp_left :
+ ∀P.(∀e:exadecim. P e = true) → (forall_exadecim (λe.P e) = true).
+ intro;
+ elim P 0;
+ [ simplify; intro; apply forall_exadecim_eqProp_left_aux; apply H; |
+   simplify; intro; apply forall_exadecim_eqProp_left_aux; apply H; |
+   autobatch; ].
+qed.
+
+lemma forall_exadecim_eqProp_right_aux :
+ ∀P.((P x0 ⊗ P x1 ⊗ P x2 ⊗ P x3 ⊗ P x4 ⊗ P x5 ⊗ P x6 ⊗ P x7 ⊗
+      P x8 ⊗ P x9 ⊗ P xA ⊗ P xB ⊗ P xC ⊗ P xD ⊗ P xE ⊗ P xF) = true) →
+    (∀e:exadecim.P e = true).
+ elim daemon.
+qed.
+
+lemma forall_exadecim_eqProp_right :
+ ∀P.(forall_exadecim (λe.P e) = true) → (∀e:exadecim.P e = true).
+ intro;
+ elim P 0;
+ [ simplify; intro; apply forall_exadecim_eqProp_right_aux; apply H; |
+   simplify; intro; apply forall_exadecim_eqProp_right_aux; apply H; |
+   autobatch; ].
+qed.
+
+lemma lt_nat_of_exadecim_16_forall: ∀e.ltb (nat_of_exadecim e) 16 = true.
+ apply forall_exadecim_eqProp_right;
+ reflexivity;
+qed.
+*)
+
+(* FINE DELL'ESPERIMENTO *)
+
+lemma lt_nat_of_exadecim_16: ∀e.nat_of_exadecim e < 16.
+ intro;
+ elim e;
+ simplify;
+ autobatch depth=17.
+qed.
+
+lemma exadecim_of_nat_mod:
+ ∀n.exadecim_of_nat n = exadecim_of_nat (n \mod 16).
+ intros;
+ apply (nat_elim1 n); intro;
+ cases m; [ intro; reflexivity | ];
+ cases n1; [ intro; reflexivity | ];
+ cases n2; [ intro; reflexivity | ];
+ cases n3; [ intro; reflexivity | ];
+ cases n4; [ intro; reflexivity | ];
+ cases n5; [ intro; reflexivity | ];
+ cases n6; [ intro; reflexivity | ];
+ cases n7; [ intro; reflexivity | ];
+ cases n8; [ intro; reflexivity | ];
+ cases n9; [ intro; reflexivity | ];
+ cases n10; [ intro; reflexivity | ];
+ cases n11; [ intro; reflexivity | ];
+ cases n12; [ intro; reflexivity | ];
+ cases n13; [ intro; reflexivity | ];
+ cases n14; [ intro; reflexivity | ];
+ cases n15; [ intro; reflexivity | ];
+ intros;
+ change in ⊢ (? ? % ?) with (exadecim_of_nat n16);
+ change in ⊢ (? ? ? (? (? % ?))) with (16 + n16);
+ rewrite > mod_plus;
+ change in ⊢ (? ? ? (? (? % ?))) with (n16 \mod 16);
+ rewrite < mod_mod;
+  [ apply H;
+    unfold lt;
+    autobatch.
+  | autobatch
+  ]
+qed.
+
+lemma nat_of_exadecim_exadecim_of_nat:
+ ∀n. nat_of_exadecim (exadecim_of_nat n) = n \mod 16.
+ intro;
+ rewrite > exadecim_of_nat_mod;
+ generalize in match (lt_mod_m_m n 16 ?); [2: autobatch | ]
+ generalize in match (n \mod 16); intro;
+ cases n1; [ intro; reflexivity | ];
+ cases n2; [ intro; reflexivity | ];
+ cases n3; [ intro; reflexivity | ];
+ cases n4; [ intro; reflexivity | ];
+ cases n5; [ intro; reflexivity | ]; 
+ cases n6; [ intro; reflexivity | ]; 
+ cases n7; [ intro; reflexivity | ];
+ cases n8; [ intro; reflexivity | ];
+ cases n9; [ intro; reflexivity | ];
+ cases n10; [ intro; reflexivity | ];
+ cases n11 [ intro; reflexivity | ];
+ cases n12; [ intro; reflexivity | ];
+ cases n13; [ intro; reflexivity | ];
+ cases n14; [ intro; reflexivity | ];
+ cases n15; [ intro; reflexivity | ];
+ cases n16; [ intro; reflexivity | ];
+ intro;
+ unfold lt in H;
+ cut False;
+  [ elim Hcut
+  | autobatch
+  ]
+qed.
+
+lemma exadecim_of_nat_nat_of_exadecim:
+ ∀e.exadecim_of_nat (nat_of_exadecim e) = e.
+ intro;
+ elim e;
+ reflexivity.
+qed.
+
+lemma plusex_ok:
+ ∀b1,b2,c.
+  match plus_ex b1 b2 c with
+   [ pairT r c' ⇒ b1 + b2 + nat_of_bool c = nat_of_exadecim r + nat_of_bool c' * 16 ].
+ intros;
+ elim b1; (elim b2; (elim c; normalize; reflexivity)).
+qed.
+
+lemma eq_eqex_S_x0_false:
+ ∀n. n < 15 → eq_ex x0 (exadecim_of_nat (S n)) = false.
+ intro;
+ cases n 0; [ intro; simplify; reflexivity | clear n];
+ cases n1 0; [ intro; simplify; reflexivity | clear n1];
+ cases n 0; [ intro; simplify; reflexivity | clear n];
+ cases n1 0; [ intro; simplify; reflexivity | clear n1];
+ cases n 0; [ intro; simplify; reflexivity | clear n];
+ cases n1 0; [ intro; simplify; reflexivity | clear n1];
+ cases n 0; [ intro; simplify; reflexivity | clear n];
+ cases n1 0; [ intro; simplify; reflexivity | clear n1];
+ cases n 0; [ intro; simplify; reflexivity | clear n];
+ cases n1 0; [ intro; simplify; reflexivity | clear n1];
+ cases n 0; [ intro; simplify; reflexivity | clear n];
+ cases n1 0; [ intro; simplify; reflexivity | clear n1];
+ cases n 0; [ intro; simplify; reflexivity | clear n];
+ cases n1 0; [ intro; simplify; reflexivity | clear n1];
+ cases n 0; [ intro; simplify; reflexivity | clear n];
+ intro;
+ unfold lt in H;
+ cut (S n1 ≤ 0);
+  [ elim (not_le_Sn_O ? Hcut)
+  | do 15 (apply le_S_S_to_le);
+    assumption
+  ]
+qed.
+
+lemma eqex_true_to_eq: ∀b,b'. eq_ex b b' = true → b=b'.
+ intros 2;
+ elim b 0;
+ elim b' 0;
+ simplify;
+ intro;
+ first [ reflexivity | destruct H ].
+qed.
+
+lemma eqex_false_to_not_eq: ∀b,b'. eq_ex b b' = false → b ≠ b'.
+ intros 2;
+ elim b 0;
+ elim b' 0;
+ simplify;
+ intro;
+ try (destruct H);
+ intro K;
+ destruct K.
+qed.
+
+(* nuovi *)
+
+lemma ok_lt_ex : ∀e1,e2:exadecim.
+ lt_ex e1 e2 = ltb e1 e2.
+ intros;
+ elim e1;
+ elim e2;
+ simplify;
+ reflexivity.
+ qed.
+
+lemma ok_le_ex : ∀e1,e2:exadecim.
+ le_ex e1 e2 = leb e1 e2.
+ intros;
+ elim e1;
+ elim e2;
+ simplify;
+ reflexivity.
+ qed.
+
+lemma ok_gt_ex : ∀e1,e2:exadecim.
+ gt_ex e1 e2 = notb (leb e1 e2).
+ intros;
+ elim e1;
+ elim e2;
+ simplify;
+ reflexivity.
+ qed.
+
+lemma ok_ge_ex : ∀e1,e2:exadecim.
+ ge_ex e1 e2 = notb (ltb e1 e2).
+ intros;
+ elim e1;
+ elim e2;
+ simplify;
+ reflexivity.
+ qed.
+
+lemma ok_pred_ex : ∀e:exadecim.
+ pred_ex e = plus_exnc e xF.
+ intros;
+ elim e;
+ simplify;
+ reflexivity.
+ qed.
+
+lemma ok_succ_ex : ∀e:exadecim.
+ succ_ex e = plus_exnc e x1.
+ intros;
+ elim e;
+ simplify;
+ reflexivity.
+ qed.
+
+lemma ok_rcr_ex : ∀e:exadecim.∀b:bool.
+ rcr_ex e b = 
+ pairT exadecim bool
+  (exadecim_of_nat ((e/2)+match b with [ true ⇒ 8 | false ⇒ 0]))
+  (eqb (mod e 2) 1).
+ intros;
+ elim e;
+ elim b;
+ simplify in ⊢ (? ? ? (? ? ? (? (? ? %)) ?));
+ simplify in ⊢ (? ? ? (? ? ? ? %));
+ simplify in ⊢ (? ? ? (? ? ? % ?));
+ simplify in ⊢ (? ? % ?);
+ reflexivity;
+qed.
+
+lemma ok_rcl_ex : ∀e:exadecim.∀b:bool.
+ rcl_ex e b =
+ pairT exadecim bool
+  (exadecim_of_nat ((mod (e*2) 16)+match b with [ true ⇒ 1 | false ⇒ 0]))
+  (notb (ltb e 8)).
+ intros;
+ elim e;
+ elim b;
+ simplify in ⊢ (? ? ? (? ? ? (? (? ? %)) ?));
+ simplify in ⊢ (? ? ? (? ? ? % ?));
+ simplify in ⊢ (? ? ? (? ? ? ? (? %)));
+ simplify in ⊢ (? ? ? (? ? ? ? %));
+ simplify in ⊢ (? ? % ?);
+ reflexivity.
+ qed.
+
+lemma ok_shr_ex : ∀e:exadecim.
+ shr_ex e =
+ pairT exadecim bool
+  (exadecim_of_nat (e/2))
+  (eqb (mod e 2) 1).
+ intros;
+ elim e;
+ simplify in ⊢ (? ? ? (? ? ? % ?));
+ simplify in ⊢ (? ? ? (? ? ? ? %));
+ simplify in ⊢ (? ? % ?);
+ reflexivity.
+qed.
+
+lemma ok_shl_ex : ∀e:exadecim.
+ shl_ex e =
+ pairT exadecim bool
+  (exadecim_of_nat (mod (e*2) 16))
+  (notb (ltb e 8)).
+ intros;
+ elim e;
+ simplify in ⊢ (? ? ? (? ? ? % ?));
+ simplify in ⊢ (? ? ? (? ? ? ? (? %)));
+ simplify in ⊢ (? ? ? (? ? ? ? %));
+ simplify in ⊢ (? ? % ?);
+ reflexivity.
+qed.
+
+lemma ok_not_ex : ∀e:exadecim.
+ e + (not_ex e) = 15.
+ intros;
+ elim e;
+ simplify;
+ reflexivity.
+qed.
+
+lemma ok_compl_ex : ∀e:exadecim.
+ e + (compl_ex e) = match gt_ex e x0 with [ true ⇒ 16 | false ⇒ 0 ].
+ intros;
+ elim e;
+ simplify;
+ reflexivity.
+qed.
+
+lemma ok_MSB_ex : ∀e:exadecim.
+ MSB_ex e = notb (ltb e 8).
+ intros;
+ elim e;
+ simplify;
+ reflexivity.
+qed.
diff --git a/helm/software/matita/library/freescale/extra.ma b/helm/software/matita/library/freescale/extra.ma
new file mode 100644 (file)
index 0000000..cb8cf78
--- /dev/null
@@ -0,0 +1,201 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/extra".
+
+include "nat/div_and_mod.ma".
+include "nat/primes.ma".
+include "list/list.ma".
+include "datatypes/constructors.ma".
+include "logic/connectives.ma".
+
+(* BOOLEANI *)
+
+(* ridefinizione degli operatori booleani, per evitare l'overloading di quelli normali *)
+definition not_bool ≝
+λb:bool.match b with [ true ⇒ false | false ⇒ true ].
+
+definition and_bool ≝
+λb1,b2:bool.match b1 with
+ [ true ⇒ b2 | false ⇒ false ].
+
+definition or_bool ≝
+λb1,b2:bool.match b1 with
+ [ true ⇒ true | false ⇒ b2 ].
+
+definition xor_bool ≝
+λb1,b2:bool.match b1 with
+ [ true ⇒ not_bool b2
+ | false ⇒ b2 ].
+
+definition eq_bool ≝
+λb1,b2:bool.match b1 with
+ [ true ⇒ b2
+ | false ⇒ not_bool b2 ].
+
+(* \ominus *)
+notation "hvbox(⊖ a)" non associative with precedence 36
+ for @{ 'not_bool $a }.
+interpretation "not_bool" 'not_bool x = 
+ (cic:/matita/freescale/extra/not_bool.con x).
+
+(* \otimes *)
+notation "hvbox(a break ⊗ b)" left associative with precedence 35
+ for @{ 'and_bool $a $b }.
+interpretation "and_bool" 'and_bool x y = 
+ (cic:/matita/freescale/extra/and_bool.con x y).
+
+(* \oplus *)
+notation "hvbox(a break ⊕ b)" left associative with precedence 34
+ for @{ 'or_bool $a $b }.
+interpretation "or_bool" 'or_bool x y = 
+ (cic:/matita/freescale/extra/or_bool.con x y).
+
+(* \odot *)
+notation "hvbox(a break ⊙ b)" left associative with precedence 33
+ for @{ 'xor_bool $a $b }.
+interpretation "xor_bool" 'xor_bool x y = 
+ (cic:/matita/freescale/extra/xor_bool.con x y).
+
+(* ProdT e' gia' definito, aggiungo Prod3T e Prod4T e Prod5T *)
+
+inductive Prod3T (T1:Type) (T2:Type) (T3:Type) : Type ≝
+tripleT : T1 → T2 → T3 → Prod3T T1 T2 T3.
+
+definition fst3T ≝
+λT1.λT2.λT3.λp:Prod3T T1 T2 T3.match p with [ tripleT x _ _ ⇒ x ].
+
+definition snd3T ≝
+λT1.λT2.λT3.λp:Prod3T T1 T2 T3.match p with [ tripleT _ x _ ⇒ x ].
+
+definition thd3T ≝
+λT1.λT2.λT3.λp:Prod3T T1 T2 T3.match p with [ tripleT _ _ x ⇒ x ].
+
+inductive Prod4T (T1:Type) (T2:Type) (T3:Type) (T4:Type) : Type ≝
+quadrupleT : T1 → T2 → T3 → T4 → Prod4T T1 T2 T3 T4.
+
+definition fst4T ≝
+λT1.λT2.λT3.λT4.λp:Prod4T T1 T2 T3 T4.match p with [ quadrupleT x _ _ _ ⇒ x ].
+
+definition snd4T ≝
+λT1.λT2.λT3.λT4.λp:Prod4T T1 T2 T3 T4.match p with [ quadrupleT _ x _ _ ⇒ x ].
+
+definition thd4T ≝
+λT1.λT2.λT3.λT4.λp:Prod4T T1 T2 T3 T4.match p with [ quadrupleT _ _ x _ ⇒ x ].
+
+definition fth4T ≝
+λT1.λT2.λT3.λT4.λp:Prod4T T1 T2 T3 T4.match p with [ quadrupleT _ _ _ x ⇒ x ].
+
+inductive Prod5T (T1:Type) (T2:Type) (T3:Type) (T4:Type) (T5:Type) : Type ≝
+quintupleT : T1 → T2 → T3 → T4 → T5 → Prod5T T1 T2 T3 T4 T5.
+
+definition fst5T ≝
+λT1.λT2.λT3.λT4.λT5.λp:Prod5T T1 T2 T3 T4 T5.match p with [ quintupleT x _ _ _ _ ⇒ x ].
+
+definition snd5T ≝
+λT1.λT2.λT3.λT4.λT5.λp:Prod5T T1 T2 T3 T4 T5.match p with [ quintupleT _ x _ _ _ ⇒ x ].
+
+definition thd5T ≝
+λT1.λT2.λT3.λT4.λT5.λp:Prod5T T1 T2 T3 T4 T5.match p with [ quintupleT _ _ x _ _ ⇒ x ].
+
+definition frth5T ≝
+λT1.λT2.λT3.λT4.λT5.λp:Prod5T T1 T2 T3 T4 T5.match p with [ quintupleT _ _ _ x _ ⇒ x ].
+
+definition ffth5T ≝
+λT1.λT2.λT3.λT4.λT5.λp:Prod5T T1 T2 T3 T4 T5.match p with [ quintupleT _ _ _ _ x ⇒ x ].
+
+(* OPTIOTN MAP *)
+
+(* option map = match ... with [ None ⇒ None ? | Some .. ⇒ .. ] *)
+definition opt_map ≝
+λT1,T2:Type.λt:option T1.λf:T1 → option T2.
+ match t with [ None ⇒ None ? | Some x ⇒ (f x) ].
+
+(* ********************** *)
+(* TEOREMI/LEMMMI/ASSIOMI *)
+(* ********************** *)
+
+axiom mod_plus: ∀a,b,m. (a + b) \mod m = (a \mod m + b \mod m) \mod m.
+axiom mod_mod: ∀a,n,m. n∣m → a \mod n = a \mod n \mod m.
+axiom eq_mod_times_n_m_m_O: ∀n,m. O < m → n * m \mod m = O.
+axiom eq_mod_to_eq_plus_mod: ∀a,b,c,m. a \mod m = b \mod m → (a+c) \mod m = (b+c) \mod m.
+axiom eq_mod_times_times_mod: ∀a,b,n,m. m = a*n → (a*b) \mod m = a * (b \mod n).
+axiom divides_to_eq_mod_mod_mod: ∀a,n,m. n∣m → a \mod m \mod n = a \mod n.
+axiom le_to_le_plus_to_le : ∀a,b,c,d.b\leq d\rarr a+b\leq c+d\rarr a\leq c.
+axiom or_lt_le : ∀n,m. n < m ∨ m ≤ n.
+
+lemma le_to_lt: ∀n,m. n ≤ m → n < S m.
+ intros;
+ autobatch.
+qed.
+
+alias num (instance 0) = "natural number".
+definition nat_of_bool ≝
+ λb:bool.match b return λ_.nat with [ true ⇒ 1 | false ⇒ 0 ].
+
+theorem lt_trans: ∀x,y,z. x < y → y < z → x < z.
+ unfold lt;
+ intros;
+ autobatch.
+qed.
+
+lemma leq_m_n_to_eq_div_n_m_S: ∀n,m:nat. 0 < m → m ≤ n → ∃z. n/m = S z.
+ intros;
+ unfold div;
+ apply (ex_intro ? ? (div_aux (pred n) (n-m) (pred m)));
+ cut (∃w.m = S w);
+  [ elim Hcut;
+    rewrite > H2;
+    rewrite > H2 in H1;
+    clear Hcut; clear H2; clear H;
+    simplify;
+    unfold in ⊢ (? ? % ?);
+    cut (∃z.n = S z);
+     [ elim Hcut; clear Hcut;
+       rewrite > H in H1;
+       rewrite > H; clear m;
+       change in ⊢ (? ? % ?)  with
+        (match leb (S a1) a with
+         [ true ⇒ O
+         | false ⇒ S (div_aux a1 ((S a1) - S a) a)]);
+       cut (S a1 ≰ a);
+        [ apply (leb_elim (S a1) a);
+           [ intro;
+             elim (Hcut H2)
+           | intro;
+             simplify;
+             reflexivity
+           ]
+        | intro;
+          autobatch
+        ]
+     | elim H1; autobatch
+     ]
+  | autobatch
+  ].
+qed.
+
+axiom daemon: False.
diff --git a/helm/software/matita/library/freescale/load_write.ma b/helm/software/matita/library/freescale/load_write.ma
new file mode 100644 (file)
index 0000000..7f743cd
--- /dev/null
@@ -0,0 +1,894 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/load_write/".
+
+(*include "/media/VIRTUOSO/freescale/model.ma".*)
+include "freescale/model.ma".
+
+(* errori possibili nel fetch *)
+inductive error_type : Type ≝
+  ILL_OP: error_type
+| ILL_FETCH_AD: error_type
+| ILL_EX_AD: error_type.
+
+(* un tipo opzione ad hoc
+   - errore: interessa solo l'errore
+   - ok: interessa info e il nuovo pc
+*)
+inductive fetch_result (A:Type) : Type ≝
+  FetchERR : error_type → fetch_result A
+| FetchOK  : A → word16 → fetch_result A.
+
+(* **************************** *)
+(* FETCH E ACCESSO ALLA MEMORIA *)
+(* **************************** *)
+
+(* ausialiaria per RS08 read *)
+(* come anticipato in status, nell'RS08 ci sono 2 registri importanti
+   memory mapped, quindi bisona intercettare la lettura.
+   NB: fare molta attenzione alle note sulle combinazioni possibili perche'
+       il comportamento della memoria nell'RS08 e' strano e ci sono
+       precise condizioni che impediscono una semantica circolare dell'accesso
+       (divergenza=assenza di definizione) *)
+definition RS08_memory_filter_read_aux ≝
+λt:memory_impl.λs:any_status RS08 t.λaddr:word16.
+λT:Type.λfREG:byte8 → option T.λfMEM:aux_mem_type t → aux_chk_type t → word16 → option T.
+match s with
+ [ mk_any_status alu mem chk _ ⇒ match alu with
+  [ mk_alu_RS08 _ _ _ _ xm psm _ _ ⇒
+(* 
+   possibili accessi al registro X
+   1) addr=000F: diretto
+   2) addr=000E (X =0F): indiretto
+   3) addr=00CF (PS=00): paging
+  
+   [NB] altre combinazioni non funzionano perche' la MCU non e' un oggetto reattivo:
+   non si possono combinare due effetti contemporaneamente!
+   per esempio accesso addr=00CE (PS=00,X=0F) non puo' produrre 2 indirezioni 
+*) 
+  match eq_w16 addr 〈〈x0,x0〉:〈x0,xF〉〉 ⊕
+        (eq_w16 addr 〈〈x0,x0〉:〈x0,xE〉〉 ⊗ eq_b8 xm 〈x0,xF〉) ⊕
+        (eq_w16 addr 〈〈x0,x0〉:〈xC,xF〉〉 ⊗ eq_b8 psm 〈x0,x0〉) with
+   [ true ⇒ fREG xm
+   | false ⇒
+(* 
+   possibili accessi al registro PS
+   1) addr=001F: diretto
+   2) addr=000E (X =1F): indiretto
+   3) addr=00DF (PS=00): paging
+*)
+    match eq_w16 addr 〈〈x0,x0〉:〈x1,xF〉〉 ⊕
+         (eq_w16 addr 〈〈x0,x0〉:〈x0,xE〉〉 ⊗ eq_b8 xm 〈x1,xF〉) ⊕
+         (eq_w16 addr 〈〈x0,x0〉:〈xD,xF〉〉 ⊗ eq_b8 psm 〈x0,x0〉) with
+     [ true ⇒ fREG psm
+     | false ⇒
+(* 
+   accesso a D[X]: se accede a [00C0-00FF] e' la RAM fisica, non il paging 
+   altrimenti sarebbero 2 indirezioni
+*)
+      match eq_w16 addr 〈〈x0,x0〉:〈x0,xE〉〉 with
+       [ true ⇒ fMEM mem chk 〈〈x0,x0〉:xm〉
+       | false ⇒ 
+(* 
+   accesso al paging: [00pp pppp ppxx xxxx] con p=PS x=addr
+*)
+        match in_range addr 〈〈x0,x0〉:〈xC,x0〉〉 〈〈x0,x0〉:〈xF,xF〉〉 with
+         [ true ⇒ fMEM mem chk (or_w16 (fstT ?? (shr_w16 (fstT ?? (shr_w16 〈psm:〈x0,x0〉〉))))
+                                             (and_w16 addr 〈〈x0,x0〉:〈x3,xF〉〉))
+(*
+   accesso normale
+*)
+         | false ⇒ fMEM mem chk addr ]]]]]].
+
+(* lettura RS08 di un byte *)
+definition RS08_memory_filter_read ≝
+λt:memory_impl.λs:any_status RS08 t.λaddr:word16.
+ RS08_memory_filter_read_aux t s addr byte8
+  (λb.Some byte8 b)
+  (λm:aux_mem_type t.λc:aux_chk_type t.λa:word16.mem_read t m c a).
+
+(* lettura RS08 di un bit *)
+definition RS08_memory_filter_read_bit ≝
+λt:memory_impl.λs:any_status RS08 t.λaddr:word16.λsub:oct.
+ RS08_memory_filter_read_aux t s addr bool
+  (λb.Some bool (getn_array8T sub bool (bits_of_byte8 b)))
+  (λm:aux_mem_type t.λc:aux_chk_type t.λa:word16.mem_read_bit t m c a sub).
+
+(* in caso di RS08 si dirotta sul filtro, altrimenti si legge direttamente *)
+definition memory_filter_read ≝
+λm:mcu_type.λt:memory_impl.match m return λm:mcu_type.any_status m t → word16 → option byte8 with
+ [ HC05 ⇒ λs:any_status HC05 t.λaddr:word16.
+  mem_read t (get_mem_desc ? t s) (get_chk_desc ? t s) addr
+ | HC08 ⇒ λs:any_status HC08 t.λaddr:word16.
+  mem_read t (get_mem_desc ? t s) (get_chk_desc ? t s) addr
+ | HCS08 ⇒ λs:any_status HCS08 t.λaddr:word16.
+  mem_read t (get_mem_desc ? t s) (get_chk_desc ? t s) addr
+ | RS08 ⇒ λs:any_status RS08 t.λaddr:word16.
+  RS08_memory_filter_read t s addr
+ ].
+
+definition memory_filter_read_bit ≝
+λm:mcu_type.λt:memory_impl.match m return λm:mcu_type.any_status m t → word16 → oct → option bool with
+ [ HC05 ⇒ λs:any_status HC05 t.λaddr:word16.λsub:oct.
+  mem_read_bit t (get_mem_desc ? t s) (get_chk_desc ? t s) addr sub
+ | HC08 ⇒ λs:any_status HC08 t.λaddr:word16.λsub:oct.
+  mem_read_bit t (get_mem_desc ? t s) (get_chk_desc ? t s) addr sub
+ | HCS08 ⇒ λs:any_status HCS08 t.λaddr:word16.λsub:oct.
+  mem_read_bit t (get_mem_desc ? t s) (get_chk_desc ? t s) addr sub
+ | RS08 ⇒ λs:any_status RS08 t.λaddr:word16.λsub:oct.
+  RS08_memory_filter_read_bit t s addr sub
+ ].
+
+(* ausialiaria per RS08 write *)
+(* come anticipato in status, nell'RS08 ci sono 2 registri importanti
+   memory mapped, quindi bisona intercettare la scrittura.
+   NB: fare molta attenzione alle note sulle combinazioni possibili perche'
+       il comportamento della memoria nell'RS08 e' strano e ci sono
+       precise condizioni che impediscono una semantica circolare dell'accesso
+       (divergenza=assenza di definizione) *)
+definition RS08_memory_filter_write_aux ≝
+λt:memory_impl.λs:any_status RS08 t.λaddr:word16.
+λfREG:byte8 → byte8.λfMEM:aux_mem_type t → aux_chk_type t → word16 → option (aux_mem_type t).
+match s with 
+ [ mk_any_status alu mem chk clk ⇒ match alu with
+  [ mk_alu_RS08 acclow pc pcm spc xm psm zfl cfl ⇒
+(* 
+   possibili accessi al registro X
+   1) addr=000F: diretto
+   2) addr=000E (X =0F): indiretto
+   3) addr=00CF (PS=00): paging
+  
+   [NB] altre combinazioni non funzionano perche' la MCU non e' un oggetto reattivo:
+   non si possono combinare due effetti contemporaneamente!
+   per esempio accesso addr=00CE (PS=00,X=0F) non puo' produrre 2 indirezioni 
+*) 
+  match eq_w16 addr 〈〈x0,x0〉:〈x0,xF〉〉 ⊕
+        (eq_w16 addr 〈〈x0,x0〉:〈x0,xE〉〉 ⊗ eq_b8 xm 〈x0,xF〉) ⊕
+        (eq_w16 addr 〈〈x0,x0〉:〈xC,xF〉〉 ⊗ eq_b8 psm 〈x0,x0〉) with
+   [ true ⇒ Some ? (mk_any_status RS08 t (mk_alu_RS08 acclow pc pcm spc (fREG xm) psm zfl cfl) mem chk clk)
+   | false ⇒
+(* 
+   possibili accessi al registro PS
+   1) addr=001F: diretto
+   2) addr=000E (X =1F): indiretto
+   3) addr=00DF (PS=00): paging
+*)
+    match eq_w16 addr 〈〈x0,x0〉:〈x1,xF〉〉 ⊕
+         (eq_w16 addr 〈〈x0,x0〉:〈x0,xE〉〉 ⊗ eq_b8 xm 〈x1,xF〉) ⊕
+         (eq_w16 addr 〈〈x0,x0〉:〈xD,xF〉〉 ⊗ eq_b8 psm 〈x0,x0〉) with
+     [ true ⇒ Some ? (mk_any_status RS08 t (mk_alu_RS08 acclow pc pcm spc xm (fREG psm) zfl cfl) mem chk clk)
+     | false ⇒
+(* 
+   accesso a D[X]: se accede a [00C0-00FF] e' la RAM fisica, non il paging 
+   altrimenti sarebbero 2 indirezioni
+*)
+      match eq_w16 addr 〈〈x0,x0〉:〈x0,xE〉〉 with
+       [ true ⇒ opt_map ?? (fMEM mem chk 〈〈x0,x0〉:xm〉)
+                 (λmem'.Some ? (mk_any_status RS08 t (mk_alu_RS08 acclow pc pcm spc xm psm zfl cfl) mem' chk clk))
+                                                      
+       | false ⇒
+(* 
+   accesso al paging: [00pp pppp ppxx xxxx] con p=PS x=addr
+*)
+        match in_range addr 〈〈x0,x0〉:〈xC,x0〉〉 〈〈x0,x0〉:〈xF,xF〉〉 with
+         [ true ⇒ opt_map ?? (fMEM mem chk (or_w16 (fstT ?? (shr_w16 (fstT ?? (shr_w16 〈psm:〈x0,x0〉〉))))
+                                                   (and_w16 addr 〈〈x0,x0〉:〈x3,xF〉〉)))
+                   (λmem'.Some ? (mk_any_status RS08 t (mk_alu_RS08 acclow pc pcm spc xm psm zfl cfl) mem' chk clk))
+(*
+   accesso normale
+*)
+         | false ⇒ opt_map ?? (fMEM mem chk addr)
+                    (λmem'.Some ? (mk_any_status RS08 t (mk_alu_RS08 acclow pc pcm spc xm psm zfl cfl) mem' chk clk)) ]]]]]].
+
+(* scrittura RS08 di un byte *)
+definition RS08_memory_filter_write ≝
+λt:memory_impl.λs:any_status RS08 t.λaddr:word16.λval:byte8.
+ RS08_memory_filter_write_aux t s addr
+  (λb.val)
+  (λm:aux_mem_type t.λc:aux_chk_type t.λa:word16.mem_update t m c a val).
+
+(* scrittura RS08 di un bit *)
+definition RS08_memory_filter_write_bit ≝
+λt:memory_impl.λs:any_status RS08 t.λaddr:word16.λsub:oct.λval:bool.
+ RS08_memory_filter_write_aux t s addr
+  (λb.byte8_of_bits (setn_array8T sub bool (bits_of_byte8 b) val))
+  (λm:aux_mem_type t.λc:aux_chk_type t.λa:word16.mem_update_bit t m c a sub val).
+
+(* in caso di RS08 si dirotta sul filtro, altrimenti si scrive direttamente *)
+definition memory_filter_write ≝
+λm:mcu_type.λt:memory_impl.match m
+ return λm:mcu_type.any_status m t → word16 → byte8 → option (any_status m t) with
+ [ HC05 ⇒ λs:any_status HC05 t.λaddr:word16.λval:byte8.
+  opt_map ?? (mem_update t (get_mem_desc ? t s) (get_chk_desc ? t s) addr val)
+   (λmem.Some ? (set_mem_desc ? t s mem)) 
+ | HC08 ⇒ λs:any_status HC08 t.λaddr:word16.λval:byte8.
+  opt_map ?? (mem_update t (get_mem_desc ? t s) (get_chk_desc ? t s) addr val)
+   (λmem.Some ? (set_mem_desc ? t s mem))
+ | HCS08 ⇒ λs:any_status HCS08 t.λaddr:word16.λval:byte8.
+  opt_map ?? (mem_update t (get_mem_desc ? t s) (get_chk_desc ? t s) addr val)
+   (λmem.Some ? (set_mem_desc ? t s mem)) 
+ | RS08 ⇒ λs:any_status RS08 t.λaddr:word16.λval:byte8.
+  RS08_memory_filter_write t s addr val
+ ].
+
+definition memory_filter_write_bit ≝
+λm:mcu_type.λt:memory_impl.match m
+ return λm:mcu_type.any_status m t → word16 → oct → bool → option (any_status m t) with
+ [ HC05 ⇒ λs:any_status HC05 t.λaddr:word16.λsub:oct.λval:bool.
+  opt_map ?? (mem_update_bit t (get_mem_desc ? t s) (get_chk_desc ? t s) addr sub val)
+   (λmem.Some ? (set_mem_desc ? t s mem)) 
+ | HC08 ⇒ λs:any_status HC08 t.λaddr:word16.λsub:oct.λval:bool.
+  opt_map ?? (mem_update_bit t (get_mem_desc ? t s) (get_chk_desc ? t s) addr sub val)
+   (λmem.Some ? (set_mem_desc ? t s mem))
+ | HCS08 ⇒ λs:any_status HCS08 t.λaddr:word16.λsub:oct.λval:bool.
+  opt_map ?? (mem_update_bit t (get_mem_desc ? t s) (get_chk_desc ? t s) addr sub val)
+   (λmem.Some ? (set_mem_desc ? t s mem)) 
+ | RS08 ⇒ λs:any_status RS08 t.λaddr:word16.λsub:oct.λval:bool.
+  RS08_memory_filter_write_bit t s addr sub val
+ ].
+
+(*
+   Da utilizzarsi solo per gli aggiornamenti di PC (per il fetch),
+   NON per il caricamento degli indiretti.
+   - il caricamento degli immediati spetta al fetcher
+     (incremento progressivo di PC ciclo per ciclo, e riempimento del prefetch
+      che a questo punto DEVE poter indirizzare qualsiasi locazione puntata da PC)
+   - il caricamento degli indiretti non spetta al fetcher
+*)
+definition filtered_inc_w16 ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λw:word16.
+ get_pc_reg m t (set_pc_reg m t s (succ_w16 w)).
+
+let rec filtered_plus_w16 (m:mcu_type) (t:memory_impl) (s:any_status m t) (w:word16) (n:nat) on n ≝
+ match n with
+  [ O ⇒ w
+  | S n' ⇒ filtered_plus_w16 m t s (filtered_inc_w16 m t s w) n' ].
+
+(* 
+   errore1: non esiste traduzione ILL_OP
+   errore2: non e' riuscito a leggere ILL_FETCH_AD
+   altrimenti OK=info+new_pc
+*)
+definition fetch ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+ let pc ≝ get_pc_reg m t s in
+ let pc_next1 ≝ filtered_inc_w16 m t s pc in
+ let pc_next2 ≝ filtered_inc_w16 m t s pc_next1 in
+ match memory_filter_read m t s pc with
+  [ None ⇒ FetchERR ? ILL_FETCH_AD
+  | Some bh ⇒ match full_info_of_word16 m (Byte bh) with
+   (* non ha trovato una traduzione con 1 byte *)
+   [ None ⇒ match m with
+    (* HC05 non esistono op a 2 byte *)
+    [ HC05 ⇒ FetchERR ? ILL_OP
+    | HC08 ⇒ match eq_b8 bh 〈x9,xE〉 with
+     (* HC08 se il primo byte e' 0x9E il secondo puo' avere senso *)
+     [ true ⇒ match memory_filter_read m t s pc_next1 with
+      [ None ⇒ FetchERR ? ILL_FETCH_AD | Some bl ⇒ match full_info_of_word16 m (Word (mk_word16 bh bl)) with
+      [ None ⇒ FetchERR ? ILL_OP | Some info ⇒ FetchOK ? info pc_next2 ]]
+     (* HC08 se il primo byte non e' 0x9E il secondo non puo' avere senso *)
+     | false ⇒ FetchERR ? ILL_OP
+     ]
+    | HCS08 ⇒ match eq_b8 bh 〈x9,xE〉 with
+     (* HCS08 se il primo byte e' 0x9E il secondo puo' avere senso *)
+     [ true ⇒ match memory_filter_read m t s pc_next1 with
+      [ None ⇒ FetchERR ? ILL_FETCH_AD | Some bl ⇒ match full_info_of_word16 m (Word (mk_word16 bh bl)) with
+      [ None ⇒ FetchERR ? ILL_OP | Some info ⇒ FetchOK ? info pc_next2 ]]
+     (* HCS08 se il primo byte non e' 0x9E il secondo non puo' avere senso *)
+     | false ⇒ FetchERR ? ILL_OP
+     ]
+    (* RS08 non esistono op a 2 byte *)
+    | RS08 ⇒ FetchERR ? ILL_OP
+    ]
+   (* ha trovato una traduzione con 1 byte *)
+   | Some info ⇒ FetchOK ? info pc_next1 ]].
+
+(* ************************ *)
+(* MODALITA' INDIRIZZAMENTO *)
+(* ************************ *)
+
+(* mattoni base *)
+(* - incrementano l'indirizzo normalmente *)
+(* - incrementano PC attraverso il filtro *)
+
+(* lettura byte da addr *)
+definition loadb_from ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λaddr:word16.λcur_pc:word16.λfetched:nat.
+ opt_map ?? (memory_filter_read m t s addr)
+  (λb.Some ? (tripleT ??? s b (filtered_plus_w16 m t s cur_pc fetched))).
+
+(* lettura bit da addr *)
+definition loadbit_from ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λaddr:word16.λsub:oct.λcur_pc:word16.λfetched:nat.
+ opt_map ?? (memory_filter_read_bit m t s addr sub)
+  (λb.Some ? (tripleT ??? s b (filtered_plus_w16 m t s cur_pc fetched))).
+
+(* lettura word da addr *)
+definition loadw_from ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λaddr:word16.λcur_pc:word16.λfetched:nat.
+ opt_map ?? (memory_filter_read m t s addr)
+  (λbh.opt_map ?? (memory_filter_read m t s (succ_w16 addr))
+   (λbl.Some ? (tripleT ??? s (mk_word16 bh bl) (filtered_plus_w16 m t s cur_pc fetched)))).
+
+(* scrittura byte su addr *)
+definition writeb_to ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λaddr:word16.λcur_pc:word16.λfetched:nat.λwriteb:byte8.
+ opt_map ?? (memory_filter_write m t s addr writeb)
+  (λtmps.Some ? (pairT ?? tmps (filtered_plus_w16 m t s cur_pc fetched))).
+
+(* scrittura bit su addr *)
+definition writebit_to ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λaddr:word16.λsub:oct.λcur_pc:word16.λfetched:nat.λwriteb:bool.
+ opt_map ?? (memory_filter_write_bit m t s addr sub writeb)
+  (λtmps.Some ? (pairT ?? tmps (filtered_plus_w16 m t s cur_pc fetched))).
+
+(* scrittura word su addr *) 
+definition writew_to ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λaddr:word16.λcur_pc:word16.λfetched:nat.λwritew:word16.
+ opt_map ?? (memory_filter_write m t s addr (w16h writew))
+  (λtmps.opt_map ?? (memory_filter_write m t s (succ_w16 addr) (w16l writew))
+    (λtmps'.Some ? (pairT ?? tmps' (filtered_plus_w16 m t s cur_pc fetched)))).
+
+(* ausiliari per definire i tipi e la lettura/scrittura *)
+
+(* ausiliaria per definire il tipo di aux_load *)
+definition aux_load_typing ≝
+λm:mcu_type.λt:memory_impl.λbyteflag:bool.
+ any_status m t → word16 → word16 → nat →
+ option (Prod3T (any_status m t) match byteflag with [ true ⇒ byte8 | false ⇒ word16 ] word16).
+
+(* per non dover ramificare i vari load in byte/word *)
+definition aux_load ≝
+λm:mcu_type.λt:memory_impl.λbyteflag:bool.match byteflag return aux_load_typing m t with
+ [ true ⇒ loadb_from m t | false ⇒ loadw_from m t ].
+
+(* ausiliaria per definire il tipo di aux_write *)
+definition aux_write_typing ≝
+λm:mcu_type.λt:memory_impl.λbyteflag:bool.
+ any_status m t → word16 → word16 → nat →
+ match byteflag with [ true ⇒ byte8 | false ⇒ word16 ] →
+ option (ProdT (any_status m t) word16).
+
+(* per non dover ramificare i vari load in byte/word *)
+definition aux_write ≝
+λm:mcu_type.λt:memory_impl.λbyteflag:bool.match byteflag return aux_write_typing m t with
+ [ true ⇒ writeb_to m t | false ⇒ writew_to m t ].
+
+(* modalita' vere e proprie *)
+
+(* lettura da [curpc]: IMM1 comportamento asimmetrico, quindi non si appoggia a loadb *)
+definition mode_IMM1_load ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
+ opt_map ?? (memory_filter_read m t s cur_pc)
+  (λb.Some ? (tripleT ??? s b (filtered_inc_w16 m t s cur_pc))).
+
+(* lettura da [curpc]: IMM2 comportamento asimmetrico, quindi non si appoggia a loadw *)
+definition mode_IMM2_load ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
+ opt_map ?? (memory_filter_read m t s cur_pc)
+  (λbh.opt_map ?? (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
+   (λbl.Some ? (tripleT ??? s (mk_word16 bh bl) (filtered_plus_w16 m t s cur_pc 2)))).
+
+(* lettura da [byte [curpc]]: true=DIR1 loadb, false=DIR1 loadw *)
+definition mode_DIR1_load ≝
+λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
+ opt_map ?? (memory_filter_read m t s cur_pc)
+  (λaddr.(aux_load m t byteflag) s 〈〈x0,x0〉:addr〉 cur_pc 1).
+
+(* lettura da [byte [curpc]]: loadbit *)
+definition mode_DIR1n_load ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.λsub:oct.
+ opt_map ?? (memory_filter_read m t s cur_pc)
+  (λaddr.loadbit_from m t  s 〈〈x0,x0〉:addr〉 sub cur_pc 1).
+
+(* scrittura su [byte [curpc]]: true=DIR1 writeb, false=DIR1 writew *)
+definition mode_DIR1_write ≝
+λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
+λwritebw:match byteflag with [ true ⇒ byte8 | false ⇒ word16 ].
+ opt_map ?? (memory_filter_read m t s cur_pc)
+  (λaddr.(aux_write m t byteflag) s 〈〈x0,x0〉:addr〉 cur_pc 1 writebw).
+
+(* scrittura su [byte [curpc]]: writebit *)
+definition mode_DIR1n_write ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.λsub:oct.λwriteb:bool.
+ opt_map ?? (memory_filter_read m t s cur_pc)
+  (λaddr.writebit_to m t s 〈〈x0,x0〉:addr〉 sub cur_pc 1 writeb).
+
+(* lettura da [word [curpc]]: true=DIR2 loadb, false=DIR2 loadw *)
+definition mode_DIR2_load ≝
+λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
+ opt_map ?? (memory_filter_read m t s cur_pc)
+  (λaddrh.opt_map ?? (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
+   (λaddrl.(aux_load m t byteflag) s 〈addrh:addrl〉 cur_pc 2)).
+
+(* scrittura su [word [curpc]]: true=DIR2 writeb, false=DIR2 writew *)
+definition mode_DIR2_write ≝
+λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
+λwritebw:match byteflag with [ true ⇒ byte8 | false ⇒ word16 ].
+ opt_map ?? (memory_filter_read m t s cur_pc)
+  (λaddrh.opt_map ?? (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
+   (λaddrl.(aux_write m t byteflag) s 〈addrh:addrl〉 cur_pc 2 writebw)).
+
+definition get_IX ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+ match m with
+  [ HC05 ⇒ opt_map ?? (get_indX_8_low_reg m t s) (λindx.Some ? 〈〈x0,x0〉:indx〉) 
+  | HC08 ⇒ opt_map ?? (get_indX_16_reg m t s) (λindx.Some ? indx) 
+  | HCS08 ⇒ opt_map ?? (get_indX_16_reg m t s) (λindx.Some ? indx) 
+  | RS08 ⇒ None ? ].
+
+(* lettura da [IX]: true=IX0 loadb, false=IX0 loadw *)
+definition mode_IX0_load ≝
+λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
+ opt_map ?? (get_IX m t s)
+  (λaddr.(aux_load m t byteflag) s addr cur_pc 0).
+
+(* scrittura su [IX]: true=IX0 writeb, false=IX0 writew *)
+definition mode_IX0_write ≝
+λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
+λwritebw:match byteflag with [ true ⇒ byte8 | false ⇒ word16 ].
+ opt_map ?? (get_IX m t s)
+  (λaddr.(aux_write m t byteflag) s addr cur_pc 0 writebw).
+
+(* lettura da [IX+byte [pc]]: true=IX1 loadb, false=IX1 loadw *)
+definition mode_IX1_load ≝
+λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
+ opt_map ?? (get_IX m t s)
+  (λaddr.opt_map ?? (memory_filter_read m t s cur_pc)
+   (λoffs.(aux_load m t byteflag) s (plus_w16nc addr 〈〈x0,x0〉:offs〉) cur_pc 1)).
+
+(* scrittura su [IX+byte [pc]]: true=IX1 writeb, false=IX1 writew *)
+definition mode_IX1_write ≝
+λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
+λwritebw:match byteflag with [ true ⇒ byte8 | false ⇒ word16 ].
+ opt_map ?? (get_IX m t s)
+  (λaddr.opt_map ?? (memory_filter_read m t s cur_pc)
+   (λoffs.(aux_write m t byteflag) s (plus_w16nc addr 〈〈x0,x0〉:offs〉) cur_pc 1 writebw)).
+
+(* lettura da [IX+word [pc]]: true=IX2 loadb, false=IX2 loadw *)
+definition mode_IX2_load ≝
+λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
+ opt_map ?? (get_IX m t s)
+  (λaddr.opt_map ?? (memory_filter_read m t s cur_pc)
+   (λoffsh.opt_map ?? (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
+    (λoffsl.(aux_load m t byteflag) s (plus_w16nc addr 〈offsh:offsl〉) cur_pc 2))).
+
+(* scrittura su [IX+word [pc]]: true=IX2 writeb, false=IX2 writew *)
+definition mode_IX2_write ≝
+λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
+λwritebw:match byteflag with [ true ⇒ byte8 | false ⇒ word16 ].
+ opt_map ?? (get_IX m t s)
+  (λaddr.opt_map ?? (memory_filter_read m t s cur_pc)
+   (λoffsh.opt_map ?? (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
+    (λoffsl.(aux_write m t byteflag) s (plus_w16nc addr 〈offsh:offsl〉) cur_pc 2 writebw))).
+
+(* lettura da [SP+byte [pc]]: true=SP1 loadb, false=SP1 loadw *)
+definition mode_SP1_load ≝
+λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
+ opt_map ?? (get_sp_reg m t s)
+  (λaddr.opt_map ?? (memory_filter_read m t s cur_pc)
+   (λoffs.(aux_load m t byteflag) s (plus_w16nc addr 〈〈x0,x0〉:offs〉) cur_pc 1)).
+
+(* scrittura su [SP+byte [pc]]: true=SP1 writeb, false=SP1 writew *)
+definition mode_SP1_write ≝
+λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
+λwritebw:match byteflag with [ true ⇒ byte8 | false ⇒ word16 ].
+ opt_map ?? (get_sp_reg m t s)
+  (λaddr.opt_map ?? (memory_filter_read m t s cur_pc)
+   (λoffs.(aux_write m t byteflag) s (plus_w16nc addr 〈〈x0,x0〉:offs〉) cur_pc 1 writebw)).
+
+(* lettura da [SP+word [pc]]: true=SP2 loadb, false=SP2 loadw *)
+definition mode_SP2_load ≝
+λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
+ opt_map ?? (get_sp_reg m t s)
+  (λaddr.opt_map ?? (memory_filter_read m t s cur_pc)
+   (λoffsh.opt_map ?? (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
+    (λoffsl.(aux_load m t byteflag) s (plus_w16nc addr 〈offsh:offsl〉) cur_pc 2))).
+
+(* scrittura su [SP+word [pc]]: true=SP2 writeb, false=SP2 writew *)
+definition mode_SP2_write ≝
+λbyteflag:bool.λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.
+λwritebw:match byteflag with [ true ⇒ byte8 | false ⇒ word16 ].
+ opt_map ?? (get_sp_reg m t s)
+  (λaddr.opt_map ?? (memory_filter_read m t s cur_pc)
+   (λoffsh.opt_map ?? (memory_filter_read m t s (filtered_inc_w16 m t s cur_pc))
+    (λoffsl.(aux_write m t byteflag) s (plus_w16nc addr 〈offsh:offsl〉) cur_pc 2 writebw))).
+
+(* ************************************** *)
+(* raccordo di tutte le possibili letture *)
+(* ************************************** *)
+
+(* H:X++ *)
+definition aux_inc_indX_16 ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+ opt_map ?? (get_indX_16_reg m t s)
+  (λX_op.opt_map ?? (set_indX_16_reg m t s (succ_w16 X_op))
+   (λs_tmp.Some ? s_tmp)).
+
+(* tutte le modalita' di lettura: false=loadb true=loadw *)
+definition multi_mode_load ≝
+λbyteflag:bool.λm:mcu_type.λt:memory_impl.
+ match byteflag
+ return λbyteflag:bool.any_status m t → word16 → instr_mode → 
+  option (Prod3T (any_status m t) match byteflag with [ true ⇒ byte8 | false ⇒ word16 ] word16)
+ with
+  (* lettura di un byte *)
+  [ true ⇒ λs:any_status m t.λcur_pc:word16.λi:instr_mode.match i with
+(* NO: non ci sono indicazioni *)
+   [ MODE_INH  ⇒ None ?
+(* restituisce A *)
+   | MODE_INHA ⇒ Some ? (tripleT ??? s (get_acc_8_low_reg m t s) cur_pc)
+(* restituisce X *)
+   | MODE_INHX ⇒ opt_map ?? (get_indX_8_low_reg m t s)
+                  (λindx.Some ? (tripleT ??? s indx cur_pc))
+(* restituisce H *)
+   | MODE_INHH ⇒ opt_map ?? (get_indX_8_high_reg m t s)
+                  (λindx.Some ? (tripleT ??? s indx cur_pc))
+
+(* preleva 1 byte immediato *) 
+   | MODE_IMM1 ⇒ mode_IMM1_load m t s cur_pc
+(* NO: solo lettura word *)
+   | MODE_IMM2 ⇒ None ?
+(* preleva 1 byte da indirizzo diretto 1 byte *) 
+   | MODE_DIR1 ⇒ mode_DIR1_load true m t s cur_pc
+(* preleva 1 byte da indirizzo diretto 1 word *)
+   | MODE_DIR2 ⇒ mode_DIR2_load true m t s cur_pc
+(* preleva 1 byte da H:X *)
+   | MODE_IX0  ⇒ mode_IX0_load true m t s cur_pc
+(* preleva 1 byte da H:X+1 byte offset *)
+   | MODE_IX1  ⇒ mode_IX1_load true m t s cur_pc
+(* preleva 1 byte da H:X+1 word offset *)
+   | MODE_IX2  ⇒ mode_IX2_load true m t s cur_pc
+(* preleva 1 byte da SP+1 byte offset *)
+   | MODE_SP1  ⇒ mode_SP1_load true m t s cur_pc
+(* preleva 1 byte da SP+1 word offset *)
+   | MODE_SP2  ⇒ mode_SP2_load true m t s cur_pc
+
+(* come DIR1, chiamare scrittura per passo2: scrittura su DIR1 *)
+   | MODE_DIR1_to_DIR1 ⇒ mode_DIR1_load true m t s cur_pc
+(* come IMM1, chiamare scrittura per passo2: scrittura su DIR1 *)
+   | MODE_IMM1_to_DIR1 ⇒ mode_IMM1_load m t s cur_pc
+(* come IX0, chiamare scrittura per passo2: scrittura su DIR1 e X++ *)
+   | MODE_IX0p_to_DIR1 ⇒ mode_IX0_load true m t s cur_pc
+(* come DIR1, chiamare scrittura per passo2: scrittura su IX0 e X++ *)
+   | MODE_DIR1_to_IX0p ⇒ mode_DIR1_load true m t s cur_pc
+
+(* NO: solo lettura word/scrittura byte *)
+   | MODE_INHA_and_IMM1 ⇒ None ?
+(* NO: solo lettura word/scrittura byte *)     
+   | MODE_INHX_and_IMM1 ⇒ None ?
+(* NO: solo lettura word *) 
+   | MODE_IMM1_and_IMM1 ⇒ None ?
+(* NO: solo lettura word/scrittura byte *) 
+   | MODE_DIR1_and_IMM1 ⇒ None ?
+(* NO: solo lettura word/scrittura byte *) 
+   | MODE_IX0_and_IMM1  ⇒ None ?
+(* NO: solo lettura word *) 
+   | MODE_IX0p_and_IMM1 ⇒ None ?
+(* NO: solo lettura word/scrittura byte *) 
+   | MODE_IX1_and_IMM1  ⇒ None ?
+(* NO: solo lettura word *) 
+   | MODE_IX1p_and_IMM1 ⇒ None ?
+(* NO: solo lettura word/scrittura byte *) 
+   | MODE_SP1_and_IMM1  ⇒ None ?
+
+(* NO: solo scrittura byte *)
+   | MODE_DIRn _          ⇒ None ?
+(* NO: solo lettura word *)
+   | MODE_DIRn_and_IMM1 _ ⇒ None ?
+(* preleva 1 byte da 0000 0000 0000 xxxxb *)
+   | MODE_TNY e           ⇒ opt_map ?? (memory_filter_read m t s 〈〈x0,x0〉:〈x0,e〉〉)
+                             (λb.Some ? (tripleT ??? s b cur_pc))
+(* preleva 1 byte da 0000 0000 000x xxxxb *)
+   | MODE_SRT e           ⇒ opt_map ?? (memory_filter_read m t s 〈〈x0,x0〉:(byte8_of_bitrigesim e)〉)
+                             (λb.Some ? (tripleT ??? s b cur_pc))
+   ]
+(* lettura di una word *)
+  | false ⇒ λs:any_status m t.λcur_pc:word16.λi:instr_mode.match i with
+(* NO: non ci sono indicazioni *)
+   [ MODE_INH  ⇒ None ?
+(* NO: solo lettura/scrittura byte *)
+   | MODE_INHA ⇒ None ?
+(* NO: solo lettura/scrittura byte *)
+   | MODE_INHX ⇒ None ?
+(* NO: solo lettura/scrittura byte *)
+   | MODE_INHH ⇒ None ?
+
+(* NO: solo lettura byte *)
+   | MODE_IMM1 ⇒ None ?
+(* preleva 1 word immediato *) 
+   | MODE_IMM2 ⇒ mode_IMM2_load m t s cur_pc
+(* preleva 1 word da indirizzo diretto 1 byte *) 
+   | MODE_DIR1 ⇒ mode_DIR1_load false m t s cur_pc
+(* preleva 1 word da indirizzo diretto 1 word *) 
+   | MODE_DIR2 ⇒ mode_DIR2_load false m t s cur_pc
+(* preleva 1 word da H:X *)
+   | MODE_IX0  ⇒ mode_IX0_load false m t s cur_pc
+(* preleva 1 word da H:X+1 byte offset *)
+   | MODE_IX1  ⇒ mode_IX1_load false m t s cur_pc
+(* preleva 1 word da H:X+1 word offset *)
+   | MODE_IX2  ⇒ mode_IX2_load false m t s cur_pc
+(* preleva 1 word da SP+1 byte offset *)
+   | MODE_SP1  ⇒ mode_SP1_load false m t s cur_pc
+(* preleva 1 word da SP+1 word offset *)
+   | MODE_SP2  ⇒ mode_SP2_load false m t s cur_pc
+
+(* NO: solo lettura/scrittura byte *)
+   | MODE_DIR1_to_DIR1 ⇒ None ?
+(* NO: solo lettura/scrittura byte *)
+   | MODE_IMM1_to_DIR1 ⇒ None ?
+(* NO: solo lettura/scrittura byte *)
+   | MODE_IX0p_to_DIR1 ⇒ None ?
+(* NO: solo lettura/scrittura byte *)
+   | MODE_DIR1_to_IX0p ⇒ None ?
+
+(* preleva 2 byte, possibilita' modificare Io argomento *)
+   | MODE_INHA_and_IMM1 ⇒ opt_map ?? (mode_IMM1_load m t s cur_pc)
+                           (λS_immb_and_PC.match S_immb_and_PC with
+                            [ tripleT _ immb cur_pc' ⇒
+                             Some ? (tripleT ??? s 〈(get_acc_8_low_reg m t s):immb〉 cur_pc')])
+(* preleva 2 byte, possibilita' modificare Io argomento *)
+   | MODE_INHX_and_IMM1 ⇒ opt_map ?? (get_indX_8_low_reg m t s)
+                           (λX_op.opt_map ?? (mode_IMM1_load m t s cur_pc)
+                            (λS_immb_and_PC.match S_immb_and_PC with
+                             [ tripleT _ immb cur_pc' ⇒
+                              Some ? (tripleT ??? s 〈X_op:immb〉 cur_pc')]))
+(* preleva 2 byte, NO possibilita' modificare Io argomento *)               
+   | MODE_IMM1_and_IMM1 ⇒ opt_map ?? (mode_IMM1_load m t s cur_pc)
+                           (λS_immb1_and_PC.match S_immb1_and_PC with
+                            [ tripleT _ immb1 cur_pc' ⇒
+                             opt_map ?? (mode_IMM1_load m t s cur_pc')
+                              (λS_immb2_and_PC.match S_immb2_and_PC with
+                               [ tripleT _ immb2 cur_pc'' ⇒
+                                Some ? (tripleT ??? s 〈immb1:immb2〉 cur_pc'')])])
+(* preleva 2 byte, possibilita' modificare Io argomento *)
+   | MODE_DIR1_and_IMM1 ⇒ opt_map ?? (mode_DIR1_load true m t s cur_pc)
+                           (λS_dirb_and_PC.match S_dirb_and_PC with
+                            [ tripleT _ dirb cur_pc' ⇒
+                             opt_map ?? (mode_IMM1_load m t s cur_pc')
+                              (λS_immb_and_PC.match S_immb_and_PC with
+                               [ tripleT _ immb cur_pc'' ⇒
+                                Some ? (tripleT ??? s 〈dirb:immb〉 cur_pc'')])])
+(* preleva 2 byte, possibilita' modificare Io argomento *)
+   | MODE_IX0_and_IMM1  ⇒ opt_map ?? (mode_IX0_load true m t s cur_pc)
+                           (λS_ixb_and_PC.match S_ixb_and_PC with
+                            [ tripleT _ ixb cur_pc' ⇒
+                             opt_map ?? (mode_IMM1_load m t s cur_pc')
+                              (λS_immb_and_PC.match S_immb_and_PC with
+                               [ tripleT _ immb cur_pc'' ⇒
+                                Some ? (tripleT ??? s 〈ixb:immb〉 cur_pc'')])])
+(* preleva 2 byte, H:X++, NO possibilita' modificare Io argomento *)
+   | MODE_IX0p_and_IMM1 ⇒ opt_map ?? (mode_IX0_load true m t s cur_pc)
+                           (λS_ixb_and_PC.match S_ixb_and_PC with
+                            [ tripleT _ ixb cur_pc' ⇒
+                             opt_map ?? (mode_IMM1_load m t s cur_pc')
+                              (λS_immb_and_PC.match S_immb_and_PC with
+                               [ tripleT _ immb cur_pc'' ⇒
+                                (* H:X++ *)
+                                opt_map ?? (aux_inc_indX_16 m t s)
+                                 (λs'.Some ? (tripleT ??? s' 〈ixb:immb〉 cur_pc''))])])
+(* preleva 2 byte, possibilita' modificare Io argomento *)
+   | MODE_IX1_and_IMM1  ⇒ opt_map ?? (mode_IX1_load true m t s cur_pc)
+                           (λS_ixb_and_PC.match S_ixb_and_PC with
+                            [ tripleT _ ixb cur_pc' ⇒
+                             opt_map ?? (mode_IMM1_load m t s cur_pc')
+                              (λS_immb_and_PC.match S_immb_and_PC with
+                               [ tripleT _ immb cur_pc'' ⇒
+                                Some ? (tripleT ??? s 〈ixb:immb〉 cur_pc'')])])
+(* preleva 2 byte, H:X++, NO possibilita' modificare Io argomento *)
+   | MODE_IX1p_and_IMM1 ⇒ opt_map ?? (mode_IX1_load true m t s cur_pc)
+                           (λS_ixb_and_PC.match S_ixb_and_PC with
+                            [ tripleT _ ixb cur_pc' ⇒
+                             opt_map ?? (mode_IMM1_load m t s cur_pc')
+                              (λS_immb_and_PC.match S_immb_and_PC with
+                               [ tripleT _ immb cur_pc'' ⇒
+                                (* H:X++ *)
+                                opt_map ?? (aux_inc_indX_16 m t s)
+                                 (λs'.Some ? (tripleT ??? s' 〈ixb:immb〉 cur_pc''))])])
+(* preleva 2 byte, possibilita' modificare Io argomento *)
+   | MODE_SP1_and_IMM1  ⇒ opt_map ?? (mode_SP1_load true m t s cur_pc)
+                           (λS_spb_and_PC.match S_spb_and_PC with
+                            [ tripleT _ spb cur_pc' ⇒
+                             opt_map ?? (mode_IMM1_load m t s cur_pc')
+                              (λS_immb_and_PC.match S_immb_and_PC with
+                               [ tripleT _ immb cur_pc'' ⇒
+                                Some ? (tripleT ??? s 〈spb:immb〉 cur_pc'')])])
+
+(* NO: solo scrittura byte *)
+   | MODE_DIRn _            ⇒ None ?
+(* preleva 2 byte, il primo e' filtrato per azzerare tutti i bit tranne n-simo *)
+   | MODE_DIRn_and_IMM1 msk ⇒ opt_map ?? (mode_DIR1n_load m t s cur_pc msk)
+                               (λS_dirbn_and_PC.match S_dirbn_and_PC with
+                                [ tripleT _ dirbn cur_pc'   ⇒
+                                 opt_map ?? (mode_IMM1_load m t s cur_pc')
+                                  (λS_immb_and_PC.match S_immb_and_PC with
+                                   [ tripleT _ immb cur_pc'' ⇒
+                                     Some ? (tripleT ??? s 〈〈x0,match dirbn with [ true ⇒ x1 | false ⇒ x0 ]〉:immb〉 cur_pc'') ])])
+(* NO: solo lettura/scrittura byte *)
+   | MODE_TNY _             ⇒ None ?
+(* NO: solo lettura/scrittura byte *)
+   | MODE_SRT _             ⇒ None ?
+   ]
+  ].
+
+(* **************************************** *)
+(* raccordo di tutte le possibili scritture *)
+(* **************************************** *)
+
+(* tutte le modalita' di scrittura: true=writeb, false=writew *)
+definition multi_mode_write ≝
+λbyteflag:bool.λm:mcu_type.λt:memory_impl.match byteflag
+ return λbyteflag:bool.any_status m t → word16 → instr_mode →
+  match byteflag with [ true ⇒ byte8 | false ⇒ word16 ] →
+  option (ProdT (any_status m t) word16) with
+ (* scrittura di un byte *)
+ [ true ⇒ λs:any_status m t.λcur_pc:word16.λi:instr_mode.λwriteb:byte8.match i with
+(* NO: non ci sono indicazioni *)
+  [ MODE_INH        ⇒ None ?
+(* scrive A *)
+  | MODE_INHA       ⇒ Some ? (pairT ?? (set_acc_8_low_reg m t s writeb) cur_pc)
+(* scrive X *)
+  | MODE_INHX       ⇒ opt_map ?? (set_indX_8_low_reg m t s writeb)
+                      (λtmps.Some ? (pairT ?? tmps cur_pc)) 
+(* scrive H *)
+  | MODE_INHH       ⇒ opt_map ?? (set_indX_8_high_reg m t s writeb)
+                       (λtmps.Some ? (pairT ?? tmps cur_pc)) 
+
+(* NO: solo lettura byte *)
+  | MODE_IMM1       ⇒ None ?
+(* NO: solo lettura word *)
+  | MODE_IMM2       ⇒ None ?
+(* scrive 1 byte su indirizzo diretto 1 byte *)
+  | MODE_DIR1       ⇒ mode_DIR1_write true m t s cur_pc writeb
+(* scrive 1 byte su indirizzo diretto 1 word *)
+  | MODE_DIR2       ⇒ mode_DIR2_write true m t s cur_pc writeb
+(* scrive 1 byte su H:X *)
+  | MODE_IX0        ⇒ mode_IX0_write true m t s cur_pc writeb
+(* scrive 1 byte su H:X+1 byte offset *)
+  | MODE_IX1        ⇒ mode_IX1_write true m t s cur_pc writeb
+(* scrive 1 byte su H:X+1 word offset *)
+  | MODE_IX2        ⇒ mode_IX2_write true m t s cur_pc writeb
+(* scrive 1 byte su SP+1 byte offset *)
+  | MODE_SP1        ⇒ mode_SP1_write true m t s cur_pc writeb
+(* scrive 1 byte su SP+1 word offset *)
+  | MODE_SP2        ⇒ mode_SP2_write true m t s cur_pc writeb
+
+(* passo2: scrittura su DIR1, passo1: lettura da DIR1 *)
+  | MODE_DIR1_to_DIR1   ⇒ mode_DIR1_write true m t s cur_pc writeb
+(* passo2: scrittura su DIR1, passo1: lettura da IMM1 *)
+  | MODE_IMM1_to_DIR1   ⇒ mode_DIR1_write true m t s cur_pc writeb
+(* passo2: scrittura su DIR1 e X++, passo1: lettura da IX0 *)
+  | MODE_IX0p_to_DIR1   ⇒ opt_map ?? (mode_DIR1_write true m t s cur_pc writeb)
+                           (λS_and_PC.match S_and_PC with [ pairT S_op PC_op ⇒
+                            (* H:X++ *)
+                            opt_map ?? (aux_inc_indX_16 m t S_op)
+                             (λS_op'.Some ? (pairT ?? S_op' PC_op))])
+(* passo2: scrittura su IX0 e X++, passo1: lettura da DIR1 *)
+  | MODE_DIR1_to_IX0p   ⇒ opt_map ?? (mode_IX0_write true m t s cur_pc writeb)
+                           (λS_and_PC.match S_and_PC with [ pairT S_op PC_op ⇒
+                            (* H:X++ *)
+                            opt_map ?? (aux_inc_indX_16 m t S_op)
+                             (λS_op'.Some ? (pairT ?? S_op' PC_op))])
+
+(* dopo aver prelevato 2 byte la possibilita' modificare Io argomento = INHA *)
+  | MODE_INHA_and_IMM1 ⇒ Some ? (pairT ?? (set_acc_8_low_reg m t s writeb) cur_pc)
+(* dopo aver prelevato 2 byte la possibilita' modificare Io argomento = INHX *)  
+  | MODE_INHX_and_IMM1 ⇒ opt_map ?? (set_indX_8_low_reg m t s writeb)
+                           (λtmps.Some ? (pairT ?? tmps cur_pc)) 
+(* NO: solo lettura word *) 
+  | MODE_IMM1_and_IMM1 ⇒ None ?
+(* dopo aver prelevato 2 byte la possibilita' modificare Io argomento = DIR1 *) 
+  | MODE_DIR1_and_IMM1 ⇒ mode_DIR1_write true m t s cur_pc writeb
+(* dopo aver prelevato 2 byte la possibilita' modificare Io argomento = IX0 *)
+  | MODE_IX0_and_IMM1  ⇒ mode_IX0_write true m t s cur_pc writeb
+(* NO: solo lettura word *) 
+  | MODE_IX0p_and_IMM1 ⇒ None ?
+(* dopo aver prelevato 2 byte la possibilita' modificare Io argomento = IX1 *)
+  | MODE_IX1_and_IMM1  ⇒ mode_IX1_write true m t s cur_pc writeb
+(* NO: solo lettura word *) 
+  | MODE_IX1p_and_IMM1 ⇒ None ?
+(* dopo aver prelevato 2 byte la possibilita' modificare Io argomento = SP1 *)
+  | MODE_SP1_and_IMM1  ⇒ mode_SP1_write true m t s cur_pc writeb
+
+(* scrive 1 byte, ma la scrittura avviene solo per n-simo bit = leggi/modifica bit/scrivi *)
+  | MODE_DIRn msk ⇒ mode_DIR1n_write m t s cur_pc msk (getn_array8T msk bool (bits_of_byte8 writeb))   
+(* NO: solo lettura word *)
+  | MODE_DIRn_and_IMM1 _ ⇒ None ?
+(* scrive 1 byte su 0000 0000 0000 xxxxb *)
+  | MODE_TNY e ⇒ opt_map ?? (memory_filter_write m t s 〈〈x0,x0〉:〈x0,e〉〉 writeb)
+                   (λtmps.Some ? (pairT ?? tmps cur_pc))
+(* scrive 1 byte su 0000 0000 000x xxxxb *)
+  | MODE_SRT e ⇒ opt_map ?? (memory_filter_write m t s 〈〈x0,x0〉:(byte8_of_bitrigesim e)〉 writeb)
+                      (λtmps.Some ? (pairT ?? tmps cur_pc)) ]
+ (* scrittura di una word *)
+ | false ⇒ λs:any_status m t.λcur_pc:word16.λi:instr_mode.λwritew:word16.match i with
+(* NO: non ci sono indicazioni *)
+  [ MODE_INH        ⇒ None ?
+(* NO: solo lettura/scrittura byte *)
+  | MODE_INHA       ⇒ None ?
+(* NO: solo lettura/scrittura byte *)
+  | MODE_INHX       ⇒ None ?
+(* NO: solo lettura/scrittura byte *)
+  | MODE_INHH       ⇒ None ?
+
+(* NO: solo lettura byte *)
+  | MODE_IMM1       ⇒ None ?
+(* NO: solo lettura word *)
+  | MODE_IMM2       ⇒ None ?
+(* scrive 1 word su indirizzo diretto 1 byte *) 
+  | MODE_DIR1       ⇒ mode_DIR1_write false m t s cur_pc writew
+(* scrive 1 word su indirizzo diretto 1 word *)
+  | MODE_DIR2       ⇒ mode_DIR2_write false m t s cur_pc writew
+(* scrive 1 word su H:X *)
+  | MODE_IX0        ⇒ mode_IX0_write false m t s cur_pc writew
+(* scrive 1 word su H:X+1 byte offset *)
+  | MODE_IX1        ⇒ mode_IX1_write false m t s cur_pc writew
+(* scrive 1 word su H:X+1 word offset *)
+  | MODE_IX2        ⇒ mode_IX2_write false m t s cur_pc writew
+(* scrive 1 word su SP+1 byte offset *)
+  | MODE_SP1        ⇒ mode_SP1_write false m t s cur_pc writew
+(* scrive 1 word su SP+1 word offset *)
+  | MODE_SP2        ⇒ mode_SP2_write false m t s cur_pc writew
+
+(* NO: solo lettura/scrittura byte *)
+  | MODE_DIR1_to_DIR1 ⇒ None ?
+(* NO: solo lettura/scrittura byte *)
+  | MODE_IMM1_to_DIR1 ⇒ None ?
+(* NO: solo lettura/scrittura byte *)
+  | MODE_IX0p_to_DIR1 ⇒ None ?
+(* NO: solo lettura/scrittura byte *)
+  | MODE_DIR1_to_IX0p ⇒ None ?
+
+(* NO: solo lettura word/scrittura byte *)
+  | MODE_INHA_and_IMM1 ⇒ None ?
+(* NO: solo lettura word/scrittura byte *)     
+  | MODE_INHX_and_IMM1 ⇒ None ?
+(* NO: solo lettura word *) 
+  | MODE_IMM1_and_IMM1 ⇒ None ?
+(* NO: solo lettura word/scrittura byte *) 
+  | MODE_DIR1_and_IMM1 ⇒ None ?
+(* NO: solo lettura word/scrittura byte *) 
+  | MODE_IX0_and_IMM1  ⇒ None ?
+(* NO: solo lettura word *) 
+  | MODE_IX0p_and_IMM1 ⇒ None ?
+(* NO: solo lettura word/scrittura byte *)
+  | MODE_IX1_and_IMM1  ⇒ None ?
+(* NO: solo lettura word *) 
+  | MODE_IX1p_and_IMM1 ⇒ None ?
+(* NO: solo lettura word/scrittura byte *) 
+  | MODE_SP1_and_IMM1  ⇒ None ?
+
+(* NO: solo scrittura byte *)
+  | MODE_DIRn _          ⇒ None ?
+(* NO: solo lettura word *)
+  | MODE_DIRn_and_IMM1 _ ⇒ None ?
+(* NO: solo lettura/scrittura byte *)
+  | MODE_TNY _           ⇒ None ?
+(* NO: solo lettura/scrittura byte *)
+  | MODE_SRT _           ⇒ None ?
+  ]
+ ].
diff --git a/helm/software/matita/library/freescale/medium_tests.ma b/helm/software/matita/library/freescale/medium_tests.ma
new file mode 100644 (file)
index 0000000..41cc96a
--- /dev/null
@@ -0,0 +1,380 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/medium_tests/".
+
+(*include "/media/VIRTUOSO/freescale/medium_tests_tools.ma".*)
+include "freescale/medium_tests_tools.ma".
+
+(* ************************ *)
+(* HCS08GB60 String Reverse *)
+(* ************************ *)
+
+(* versione ridotta, in cui non si riazzerano gli elementi di counters *)
+definition dTest_HCS08_sReverse_source : word16 → (list byte8) ≝
+λelems:word16.
+let m ≝ HCS08 in source_to_byte8 m (
+(* BEFORE: A=0x00 H:X=0x0D4B SP=0x0D4A PC=0x18E0 Z=true *)
+
+(* static unsigned char dati[3072]={...};
+
+   void swap(unsigned char *a, unsigned char *b)
+    { unsigned char tmp=*a; *a=*b; *b=tmp; return; } *)
+
+(* [0x18C8]    allineamento *) (compile m ? NOP maINH I) @
+
+(* [0x18C9]       PSHX      *) (compile m ? PSHX maINH I) @
+(* [0x18CA]       PSHH      *) (compile m ? PSHH maINH I) @
+(* [0x18CB]       LDHX 5,SP *) (compile m ? LDHX (maSP1 〈x0,x5〉) I) @
+(* [0x18CE]       LDA ,X    *) (compile m ? LDA maIX0 I) @
+(* [0x18CF]       LDHX 1,SP *) (compile m ? LDHX (maSP1 〈x0,x1〉) I) @
+(* [0x18D2]       PSHA      *) (compile m ? PSHA maINH I) @
+(* [0x18D3]       LDA ,X    *) (compile m ? LDA maIX0 I) @
+(* [0x18D4]       LDHX 6,SP *) (compile m ? LDHX (maSP1 〈x0,x6〉) I) @
+(* [0x18D7]       STA ,X    *) (compile m ? STA maIX0 I) @
+(* [0x18D8]       LDHX 2,SP *) (compile m ? LDHX (maSP1 〈x0,x2〉) I) @
+(* [0x18DB]       PULA      *) (compile m ? PULA maINH I) @
+(* [0x18DC]       STA ,X    *) (compile m ? STA maIX0 I) @
+(* [0x18DD]       AIS #2    *) (compile m ? AIS (maIMM1 〈x0,x2〉) I) @
+(* [0x18DF]       RTS       *) (compile m ? RTS maINH I) @
+
+(* void main(void)
+   {
+   unsigned int pos=0,limit=0;
+   for(limit=3072;pos<(limit/2);pos++)
+    { swap(&dati[pos],&dati[limit-pos-1]); } *)
+
+(* [0x18E0]       LDHX #LUNG      *) (compile m ? LDHX (maIMM2 elems) I) @
+(* [0x18E3]       STHX 4,SP       *) (compile m ? STHX (maSP1 〈x0,x4〉) I) @
+(* [0x18E6] 20 32 BRA *+52 ; 191A *) (compile m ? BRA (maIMM1 〈x3,x2〉) I) @
+(* [0x18E8]       TSX             *) (compile m ? TSX maINH I) @
+(* [0x18E9]       LDA 2,X         *) (compile m ? LDA (maIX1 〈x0,x2〉) I) @
+(* [0x18EB]       ADD #0x00       *) (compile m ? ADD (maIMM1 〈x0,x0〉) I) @
+(* [0x18ED]       PSHA            *) (compile m ? PSHA maINH I) @
+(* [0x18EE]       LDA 1,X         *) (compile m ? LDA (maIX1 〈x0,x1〉) I) @
+(* [0x18F0]       ADC #0x01       *) (compile m ? ADC (maIMM1 〈x0,x1〉) I) @
+(* [0x18F2]       PSHA            *) (compile m ? PSHA maINH I) @
+(* [0x18F3]       LDA 4,X         *) (compile m ? LDA (maIX1 〈x0,x4〉) I) @
+(* [0x18F5]       SUB 2,X         *) (compile m ? SUB (maIX1 〈x0,x2〉) I) @
+(* [0x18F7]       STA ,X          *) (compile m ? STA maIX0 I) @
+(* [0x18F8]       LDA 3,X         *) (compile m ? LDA (maIX1 〈x0,x3〉) I) @
+(* [0x18FA]       SBC 1,X         *) (compile m ? SBC (maIX1 〈x0,x1〉) I) @
+(* [0x18FC]       PSHA            *) (compile m ? PSHA maINH I) @
+(* [0x18FD]       LDX ,X          *) (compile m ? LDX maIX0 I) @
+(* [0x18FE]       PULH            *) (compile m ? PULH maINH I) @
+(* [0x18FF]       AIX #-1         *) (compile m ? AIX (maIMM1 〈xF,xF〉) I) @
+(* [0x1901]       TXA             *) (compile m ? TXA maINH I) @
+(* [0x1902]       ADD #0x00       *) (compile m ? ADD (maIMM1 〈x0,x0〉) I) @
+(* [0x1904]       PSHH            *) (compile m ? PSHH maINH I) @
+(* [0x1905]       TSX             *) (compile m ? TSX maINH I) @
+(* [0x1906]       STA 3,X         *) (compile m ? STA (maIX1 〈x0,x3〉) I) @
+(* [0x1908]       PULA            *) (compile m ? PULA maINH I) @
+(* [0x1909]       ADC #0x01       *) (compile m ? ADC (maIMM1 〈x0,x1〉) I) @
+(* [0x190B]       LDX 3,X         *) (compile m ? LDX (maIX1 〈x0,x3〉) I) @
+(* [0x190D]       PSHA            *) (compile m ? PSHA maINH I) @
+(* [0x190E]       PULH            *) (compile m ? PULH maINH I) @
+(* [0x190F] AD B8 BSR *-70 ; 18C9 *) (compile m ? BSR (maIMM1 〈xB,x8〉) I) @
+(* [0x1911]       AIS #2          *) (compile m ? AIS (maIMM1 〈x0,x2〉) I) @
+(* [0x1913]       TSX             *) (compile m ? TSX maINH I) @
+(* [0x1914]       INC 2,X         *) (compile m ? INC (maIX1 〈x0,x2〉) I) @
+(* [0x1916] 26 02 BNE *+4 ; 191A  *) (compile m ? BNE (maIMM1 〈x0,x2〉) I) @
+(* [0x1918]       INC 1,X         *) (compile m ? INC (maIX1 〈x0,x1〉) I) @
+(* [0x191A]       TSX             *) (compile m ? TSX maINH I) @
+(* [0x191B]       LDA 3,X         *) (compile m ? LDA (maIX1 〈x0,x3〉) I) @
+(* [0x191D]       PSHA            *) (compile m ? PSHA maINH I) @
+(* [0x191E]       PULH            *) (compile m ? PULH maINH I) @
+(* [0x191F]       LSRA            *) (compile m ? LSR maINHA I) @
+(* [0x1920]       TSX             *) (compile m ? TSX maINH I) @
+(* [0x1921]       LDX 4,X         *) (compile m ? LDX (maIX1 〈x0,x4〉) I) @
+(* [0x1923]       RORX            *) (compile m ? ROR maINHX I) @
+(* [0x1924]       PSHA            *) (compile m ? PSHA maINH I) @
+(* [0x1925]       PULH            *) (compile m ? PULH maINH I) @
+(* [0x1926]       CPHX 2,SP       *) (compile m ? CPHX (maSP1 〈x0,x2〉) I) @
+(* [0x1929] 22 BD BHI *-65 ; 18E8 *) (compile m ? BHI (maIMM1 〈xB,xD〉) I)
+
+(* [0x192B] attraverso simulazione in CodeWarrior si puo' enunciare che dopo
+            42+79*n+(n>>9) ci sara' il reverse di n byte (PARI) e
+            H:X=n/2 *)
+ ).
+
+(* creazione del processore+caricamento+impostazione registri *)
+definition dTest_HCS08_sReverse_status ≝
+λt:memory_impl.
+λHX_op:word16.
+λelems:word16.
+λdata:list byte8.
+ set_z_flag HCS08 t (* Z<-true *)
+  (setweak_sp_reg HCS08 t (* SP<-0x0D4A *)
+   (setweak_indX_16_reg HCS08 t (* H:X<-HX_op *)
+    (set_pc_reg HCS08 t (* PC<-0x18E0 *)
+     (start_of_mcu_version
+      MC9S08GB60 t
+      (load_from_source_at t (* carica data in RAM:dTest_HCS08_RAM *)
+       (load_from_source_at t (zero_memory t) (* carica source in ROM:dTest_HCS08_prog *)
+         (dTest_HCS08_sReverse_source elems) dTest_HCS08_prog)
+        data dTest_HCS08_RAM)
+      (build_memory_type_of_mcu_version MC9S08GB60 t)
+      (mk_byte8 x0 x0) (mk_byte8 x0 x0) (* non deterministici tutti a 0 *)
+      false false false false false false) (* non deterministici tutti a 0 *)
+     (mk_word16 (mk_byte8 x1 x8) (mk_byte8 xE x0)))
+    HX_op)
+   (mk_word16 (mk_byte8 x0 xD) (mk_byte8 x4 xA)))
+  true.
+
+(* parametrizzazione dell'enunciato del teorema *)
+(* primo sbozzo: confronto esecuzione con hexdump... *)
+lemma dTest_HCS08_sReverse_dump_aux ≝
+λt:memory_impl.λstring:list byte8.
+ (* 1) la stringa deve avere una lunghezza ∈ [0,3072] *)
+ (byte8_bounded_strlen string 〈〈x0,xC〉:〈x0,x0〉〉) ∧
+ (* 2) la stringa deve avere lunghezza pari *)
+ ((and_b8 (w16l (byte8_strlen string)) 〈x0,x1〉) = 〈x0,x0〉) ∧
+ (* 3) match di esecuzione su tempo in forma di tempo esatto *)
+ (match execute HCS08 t
+  (* parametri IN: t,H:X,strlen(string),string *)
+  (TickOK ? (dTest_HCS08_sReverse_status t 〈〈x0,xD〉:〈x4,xB〉〉 (byte8_strlen string) string))
+  (* tempo di esecuzione 25700+150*n *)
+  (42+79*(byte8_strlen string)+((byte8_strlen string)/512)) with
+   [ TickERR s _ ⇒ None ?
+   (* azzeramento tutta RAM tranne dati *)
+   | TickSUSP s _ ⇒ None ? 
+   | TickOK s ⇒ Some ? (byte8_hexdump t (get_mem_desc HCS08 t s) dTest_HCS08_RAM (byte8_strlen string))
+   ] =
+  Some ? (byte8_reverse string)).
+
+(* primo sbozzo: confronto esecuzione con hexdump... *)
+lemma dTest_HCS08_sReverse_dump :
+ dTest_HCS08_sReverse_dump_aux MEM_TREE dTest_random_32.
+ unfold dTest_HCS08_sReverse_dump_aux;
+ split;
+ [ split; [ normalize in ⊢ (%); autobatch ] reflexivity ]
+ reflexivity.
+qed.
+
+(* parametrizzazione dell'enunciato del teorema *)
+(* dimostrazione senza svolgimento degli stati *)
+lemma dTest_HCS08_sReverse_aux ≝
+λt:memory_impl.λstring:list byte8.
+ (* 1) la stringa deve avere una lunghezza ∈ [0,3072] *)
+ (byte8_bounded_strlen string 〈〈x0,xC〉:〈x0,x0〉〉) ∧
+ (* 2) la stringa deve avere lunghezza pari *)
+ ((and_b8 (w16l (byte8_strlen string)) 〈x0,x1〉) = 〈x0,x0〉) ∧
+ (* 3) match di esecuzione su tempo in forma di tempo esatto *)
+ (match execute HCS08 t
+  (* parametri IN: t,H:X,strlen(string),string *)
+  (TickOK ? (dTest_HCS08_sReverse_status t 〈〈x0,xD〉:〈x4,xB〉〉 (byte8_strlen string) string))
+  (* tempo di esecuzione 25700+150*n *)
+  (42+79*(byte8_strlen string)+((byte8_strlen string)/512)) with
+   [ TickERR s _ ⇒ None ?
+   (* azzeramento tutta RAM tranne dati *)
+   | TickSUSP s _ ⇒ None ? 
+   | TickOK s ⇒ Some ? (set_mem_desc HCS08 t s (load_from_source_at t (get_mem_desc HCS08 t s) dTest_bytes 〈〈x0,xD〉:〈x0,x0〉〉))
+   ] =
+  Some ? (set_pc_reg HCS08 t
+   (dTest_HCS08_sReverse_status t (fstT ?? (shr_w16 (byte8_strlen string))) (byte8_strlen string) (byte8_reverse string)) 
+   (mk_word16 (mk_byte8 x1 x9) (mk_byte8 x2 xB)))).
+
+(* dimostrazione senza svolgimento degli stati *)
+(* sembra corretta, ma non basta la memoria per arrivare in fondo con quelli grandi! *)
+(* NB: sono gia' pronti dei pattern piu' impegnativi in medium_tests_tool: *)
+(*     dTest_random_32/64/128/256/512/1024/2048/3072 generati con Mathematica *)
+(*     ex: dTest_HCS08_sReverse_aux MEM_TREE dTest_random_3072 sarebbe gia' un bell'obbiettivo! *)
+lemma dTest_HCS08_sReverse :
+ dTest_HCS08_sReverse_aux MEM_TREE dTest_random_32.
+ unfold dTest_HCS08_sReverse_aux;
+ split;
+ [ split; [ normalize in ⊢ (%); autobatch ] reflexivity ]
+ reflexivity.
+qed.
+
+(* *********************** *)
+(* HCS08GB60 Counting Sort *)
+(* *********************** *)
+
+(* versione ridotta, in cui non si riazzerano gli elementi di counters *)
+definition dTest_HCS08_cSort_source : word16 → (list byte8) ≝
+λelems:word16.
+let m ≝ HCS08 in source_to_byte8 m (
+(* BEFORE: A=0x00 H:X=0x0F4C SP=0x0F4B PC=0x18C8 Z=true *)
+
+(* /* IPOTESI: INIT VARIABILI+ARRAY GIA' ESEGUITO */
+   static unsigned int counters[256]={ campitura di 0 };
+   static unsigned char dati[3072]={ dati random };
+
+   void CountingSort(void)
+    {
+    unsigned int index=0,position=0; *)
+
+(* /* TESI: CODICE DA ESEGUIRE
+
+    /* calcolo del # ripetizioni degli elementi byte */
+    for(;index<3072;index++)
+     { counters[dati[index]]++; } *)
+
+(* [0x18C8] 20 1D       BRA *+31;18E7 [3] *) (compile m ? BRA (maIMM1 〈x1,xD〉) I) @
+(* [0x18CA] 9E FE 01    LDHX 1,SP     [5] *) (compile m ? LDHX (maSP1 〈x0,x1〉) I) @
+(* [0x18CD] D6 01 00    LDA 256,X     [4] *) (compile m ? LDA (maIX2 〈〈x0,x1〉:〈x0,x0〉〉) I) @
+(* [0x18D0] 48          LSLA          [1] *) (compile m ? ASL maINHA I) @
+(* [0x18D1] 5F          CLRX          [1] *) (compile m ? CLR maINHX I) @
+(* [0x18D2] 59          ROLX          [1] *) (compile m ? ROL maINHX I) @
+(* [0x18D3] AB 00       ADD #0x00     [2] *) (compile m ? ADD (maIMM1 〈x0,x0〉) I) @
+(* [0x18D5] 87          PSHA          [2] *) (compile m ? PSHA maINH I) @
+(* [0x18D6] 9F          TXA           [1] *) (compile m ? TXA maINH I) @
+(* [0x18D7] A9 0D       ADC #0x0D     [2] *) (compile m ? ADC (maIMM1 〈x0,xD〉) I) @
+(* [0x18D9] 87          PSHA          [2] *) (compile m ? PSHA maINH I) @
+(* [0x18DA] 8A          PULH          [3] *) (compile m ? PULH maINH I) @
+(* [0x18DB] 88          PULX          [3] *) (compile m ? PULX maINH I) @
+(* [0x18DC] 6C 01       INC 1,X       [5] *) (compile m ? INC (maIX1 〈x0,x1〉) I) @
+(* [0x18DE] 26 01       BNE *+3       [3] *) (compile m ? BNE (maIMM1 〈x0,x1〉) I) @
+(* [0x18E0] 7C          INC ,X        [4] *) (compile m ? INC maIX0 I) @
+(* [0x18E1] 95          TSX           [2] *) (compile m ? TSX maINH I) @
+(* [0x18E2] 6C 01       INC 1,X       [5] *) (compile m ? INC (maIX1 〈x0,x1〉) I) @
+(* [0x18E4] 26 01       BNE *+3       [3] *) (compile m ? BNE (maIMM1 〈x0,x1〉) I) @
+(* [0x18E6] 7C          INC ,X        [4] *) (compile m ? INC maIX0 I) @
+(* [0x18E7] 9E FE 01    LDHX 1,SP     [5] *) (compile m ? LDHX (maSP1 〈x0,x1〉) I) @
+(* [0x18EA] 65 0C 00    CPHX #0x0C00  [3] *) (compile m ? CPHX (maIMM2 elems) I) @ (* dimensione dei dati al massimo 0x0C00 *)
+(* [0x18ED] 25 DB       BCS *-35;18CA [3] *) (compile m ? BCS (maIMM1 〈xD,xB〉) I) @
+
+(* /* sovrascrittura di dati per produrre la versione ordinata */
+   for(index=0;index<256;index++)
+    {
+    while(counters[index]--)
+     { dati[position++]=index; }
+    } *)
+
+(* [0x18EF] 95          TSX          *) (compile m ? TSX maINH I) @
+(* [0x18F0] 6F 01       CLR 1,X      *) (compile m ? CLR (maIX1 〈x0,x1〉) I) @
+(* [0x18F2] 7F          CLR ,X       *) (compile m ? CLR maIX0 I) @
+(* [0x18F3] 20 0E       BRA *+16     *) (compile m ? BRA (maIMM1 〈x0,xE〉) I) @
+(* [0x18F5] 95          TSX          *) (compile m ? TSX maINH I) @
+(* [0x18F6] E6 01       LDA 1,X      *) (compile m ? LDA (maIX1 〈x0,x1〉) I) @
+(* [0x18F8] 9E FE 03    LDHX 3,SP    *) (compile m ? LDHX (maSP1 〈x0,x3〉) I) @
+(* [0x18FB] D7 01 00    STA 256,X    *) (compile m ? STA (maIX2 〈〈x0,x1〉:〈x0,x0〉〉) I) @
+(* [0x18FE] AF 01       AIX #1       *) (compile m ? AIX (maIMM1 〈x0,x1〉) I) @
+(* [0x1900] 9E FF 03    STHX 3,SP    *) (compile m ? STHX (maSP1 〈x0,x3〉) I) @
+(* [0x1903] 95          TSX          *) (compile m ? TSX maINH I) @
+(* [0x1904] EE 01       LDX 1,X      *) (compile m ? LDX (maIX1 〈x0,x1〉) I) @
+(* [0x1906] 58          LSLX         *) (compile m ? ASL maINHX I) @
+(* [0x1907] 9E E6 01    LDA 1,SP     *) (compile m ? LDA (maSP1 〈x0,x1〉) I) @
+(* [0x190A] 49          ROLA         *) (compile m ? ROL maINHA I) @
+(* [0x190B] 87          PSHA         *) (compile m ? PSHA maINH I) @
+(* [0x190C] 8A          PULH         *) (compile m ? PULH maINH I) @
+(* [0x190D] 89          PSHX         *) (compile m ? PSHX maINH I) @
+(* [0x190E] 9E BE 0D 00 LDHX 3328,X  *) (compile m ? LDHX (maIX2 〈〈x0,xD〉:〈x0,x0〉〉) I) @
+(* [0x1912] 89          PSHX         *) (compile m ? PSHX maINH I) @
+(* [0x1913] 8B          PSHH         *) (compile m ? PSHH maINH I) @
+(* [0x1914] AF FF       AIX #-1      *) (compile m ? AIX (maIMM1 〈xF,xF〉) I) @
+(* [0x1916] 8B          PSHH         *) (compile m ? PSHH maINH I) @
+(* [0x1917] 87          PSHA         *) (compile m ? PSHA maINH I) @
+(* [0x1918] 8A          PULH         *) (compile m ? PULH maINH I) @
+(* [0x1919] 89          PSHX         *) (compile m ? PSHX maINH I) @
+(* [0x191A] 9E EE 05    LDX 5,SP     *) (compile m ? LDX (maSP1 〈x0,x5〉) I) @
+(* [0x191D] 86          PULA         *) (compile m ? PULA maINH I) @
+(* [0x191E] D7 0D 01    STA 3329,X   *) (compile m ? STA (maIX2 〈〈x0,xD〉:〈x0,x1〉〉) I) @
+(* [0x1921] 86          PULA         *) (compile m ? PULA maINH I) @
+(* [0x1922] D7 0D 00    STA 3328,X   *) (compile m ? STA (maIX2 〈〈x0,xD〉:〈x0,x0〉〉) I) @
+(* [0x1925] 8A          PULH         *) (compile m ? PULH maINH I) @
+(* [0x1926] 88          PULX         *) (compile m ? PULX maINH I) @
+(* [0x1927] 65 00 00    CPHX #0x0000 *) (compile m ? CPHX (maIMM2 〈〈x0,x0〉:〈x0,x0〉〉) I) @
+(* [0x192A] 8A          PULH         *) (compile m ? PULH maINH I) @
+(* [0x192B] 26 C8       BNE *-54     *) (compile m ? BNE (maIMM1 〈xC,x8〉) I) @
+(* [0x192D] 95          TSX          *) (compile m ? TSX maINH I) @
+(* [0x192E] 6C 01       INC 1,X      *) (compile m ? INC (maIX1 〈x0,x1〉) I) @
+(* [0x1930] 26 01       BNE *+3      *) (compile m ? BNE (maIMM1 〈x0,x1〉) I) @
+(* [0x1932] 7C          INC ,X       *) (compile m ? INC maIX0 I) @
+(* [0x1933] 9E FE 01    LDHX 1,SP    *) (compile m ? LDHX (maSP1 〈x0,x1〉) I) @
+(* [0x1936] 65 01 00    CPHX #0x0100 *) (compile m ? CPHX (maIMM2 〈〈x0,x1〉:〈x0,x0〉〉) I) @
+(* [0x1939] 25 C8       BNE *-54     *) (compile m ? BNE (maIMM1 〈xC,x8〉) I) @
+(* [0x193B] 8E                       *) (compile m ? STOP maINH I)
+
+(* [0x193C] attraverso simulazione in CodeWarrior si puo' enunciare che dopo
+            25700+150n si sara' entrati in stato STOP corrispondente con ordinamento
+            di n byte, A=0xFF H:X=0x0100 *)
+ ).
+
+(* creazione del processore+caricamento+impostazione registri *)
+definition dTest_HCS08_cSort_status ≝
+λt:memory_impl.
+λA_op:byte8.
+λHX_op:word16.
+λelems:word16.
+λdata:list byte8.
+ set_acc_8_low_reg HCS08 t (* A<-A_op *)
+  (set_z_flag HCS08 t (* Z<-true *)
+   (setweak_sp_reg HCS08 t (* SP<-0x0F4B *)
+    (setweak_indX_16_reg HCS08 t (* H:X<-HX_op *)
+     (set_pc_reg HCS08 t (* PC<-dTest_HCS08_prog *)
+      (start_of_mcu_version
+       MC9S08GB60 t
+       (load_from_source_at t (* carica data in RAM:dTest_HCS08_RAM *)
+        (load_from_source_at t (zero_memory t) (* carica source in ROM:dTest_HCS08_prog *)
+          (dTest_HCS08_cSort_source elems) dTest_HCS08_prog)
+         data dTest_HCS08_RAM)
+       (build_memory_type_of_mcu_version MC9S08GB60 t)
+       (mk_byte8 x0 x0) (mk_byte8 x0 x0) (* non deterministici tutti a 0 *)
+       false false false false false false) (* non deterministici tutti a 0 *)
+      dTest_HCS08_prog)
+     HX_op)
+    (mk_word16 (mk_byte8 x0 xF) (mk_byte8 x4 xB)))
+   true)
+  A_op.
+
+(* parametrizzazione dell'enunciato del teorema parziale *)
+lemma dTest_HCS08_cSort_aux ≝
+λt:memory_impl.λstring:list byte8.
+ (* 1) la stringa deve avere una lunghezza ∈ [0,3072] *)
+ (byte8_bounded_strlen string 〈〈x0,xC〉:〈x0,x0〉〉) ∧
+ (* 2) match di esecuzione su tempo in forma di upperbound *)
+ (match execute HCS08 t
+  (* parametri IN: t,A,H:X,strlen(string),string *)
+  (TickOK ? (dTest_HCS08_cSort_status t 〈x0,x0〉 〈〈x0,xF〉:〈x4,xC〉〉 (byte8_strlen string) string))
+  (* tempo di esecuzione 25700+150*n *)
+  ((nat_of_word16 〈〈x6,x4〉:〈x6,x4〉〉)+(nat_of_byte8 〈x9,x6〉)*(nat_of_word16 (byte8_strlen string))) with
+   [ TickERR s _ ⇒ None ?
+   (* azzeramento tutta RAM tranne dati *)
+   | TickSUSP s _ ⇒ Some ? (set_mem_desc HCS08 t s (load_from_source_at t (get_mem_desc HCS08 t s) dTest_bytes 〈〈x0,xD〉:〈x0,x0〉〉))
+   | TickOK s ⇒ None ?
+   ] =
+  Some ? (set_pc_reg HCS08 t
+   (dTest_HCS08_cSort_status t 〈xF,xF〉 〈〈x0,x1〉:〈x0,x0〉〉 (byte8_strlen string) (byte8_list_ordering string)) 
+   (mk_word16 (mk_byte8 x1 x9) (mk_byte8 x3 xC)))).
+
+(* dimostrazione senza svolgimento degli stati *)
+(* sembra corretta, ma non basta la memoria per arrivare in fondo nemmeno col piccolo! *)
+(* NB: sono gia' pronti dei pattern piu' impegnativi in medium_tests_tool: *)
+(*     dTest_random_32/64/128/256/512/1024/2048/3072 generati con Mathematica *)
+(*     ex: dTest_HCS08_cSort_aux MEM_TREE dTest_random_3072 sarebbe gia' un bell'obbiettivo! *)
+
+(*
+lemma dTest_HCS08_cSort :
+ dTest_HCS08_cSort_aux MEM_TREE dTest_random_32.
+ unfold dTest_HCS08_cSort_aux;
+ split;
+ [ normalize in ⊢ (%); autobatch ]
+ reflexivity.
+qed.
+*)
diff --git a/helm/software/matita/library/freescale/medium_tests_tools.ma b/helm/software/matita/library/freescale/medium_tests_tools.ma
new file mode 100644 (file)
index 0000000..0b273b0
--- /dev/null
@@ -0,0 +1,446 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/medium_tests_tools/".
+
+(*include "/media/VIRTUOSO/freescale/multivm.ma".*)
+include "freescale/multivm.ma".
+
+(* ********* *)
+(* INDIRIZZI *)
+(* ********* *)
+
+(* specifico per MC9S08GB60 in modo da caricare codice compilato da CodeWarrior *)
+(* l'obbiettivo e' dimostrare una routine scritta in C *)
+(* passo 1 e' formalizzare l'uso di 3Kb dei 4Kb di RAM [0x0100-0x0CFF] *)
+definition dTest_HCS08_RAM ≝ 〈〈x0,x1〉:〈x0,x0〉〉.
+definition dTest_HCS08_prog ≝ 〈〈x1,x8〉:〈xC,x8〉〉.
+
+(* ***** *)
+(* TOOLS *)
+(* ***** *)
+
+(* visita di un albero da 256B di elementi: ln16(256)=2 passaggi *)
+definition dTest_visit ≝
+λdata:Prod16T (Prod16T (list byte8)).λaddr:byte8.
+ getn_array16T (b8l addr) ?
+  (getn_array16T (b8h addr) ? data).
+
+(* scrittura di un elemento in un albero da 256B *)
+definition dTest_update ≝
+λdata:Prod16T (Prod16T (list byte8)).λaddr:byte8.λv:list byte8.
+ let lev2 ≝ getn_array16T (b8h addr) ? data in
+ setn_array16T (b8h addr) ? data
+  (setn_array16T (b8l addr) ? lev2 v) .
+
+(* array a 0 *)
+definition dTest_zero_array ≝
+let elem ≝ nil byte8 in
+let lev2 ≝ array_16T ? 
+           elem elem elem elem elem elem elem elem
+           elem elem elem elem elem elem elem elem
+           in
+let lev1 ≝ array_16T ?
+           lev2 lev2 lev2 lev2 lev2 lev2 lev2 lev2
+           lev2 lev2 lev2 lev2 lev2 lev2 lev2 lev2
+           in
+lev1.
+
+(* incrementa n-simo elemento *)
+definition dTest_inc ≝
+λdata:Prod16T (Prod16T (list byte8)).λaddr:byte8.
+ dTest_update data addr ((dTest_visit data addr)@[ addr ]).
+
+(* costruisce una lista a partire dai conteggi per elemento *)
+definition dTest_build_list_from_count ≝
+λdata:Prod16T (Prod16T (list byte8)).
+ let aux1 ≝ λparam1:Prod16T (list byte8).
+  match param1 with
+   [ array_16T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15 ⇒
+    e00@e01@e02@e03@e04@e05@e06@e07@e08@e09@e10@e11@e12@e13@e14@e15 ] in
+ let aux2 ≝ λparam2:Prod16T (Prod16T (list byte8)).
+  match param2 with
+   [ array_16T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15 ⇒
+    (aux1 e00)@(aux1 e01)@(aux1 e02)@(aux1 e03)@(aux1 e04)@(aux1 e05)@(aux1 e06)@(aux1 e07)@
+    (aux1 e08)@(aux1 e09)@(aux1 e10)@(aux1 e11)@(aux1 e12)@(aux1 e13)@(aux1 e14)@(aux1 e15) ] in
+ aux2 data.
+
+(* ci sono ora tutti gli elementi per definire l'ordinamento *)
+(* di una lista di byte8 secondo il counting sort *)
+let rec byte8_list_ordering_aux (source:list byte8) (count:Prod16T (Prod16T (list byte8))) on source ≝
+ match source with
+  [ nil ⇒ dTest_build_list_from_count count
+  | cons hd tl ⇒ byte8_list_ordering_aux tl (dTest_inc count hd)
+  ].
+
+definition byte8_list_ordering ≝
+λsource:list byte8.byte8_list_ordering_aux source dTest_zero_array.
+
+(* strlen esadecimale limitato, considerando overflow! *)
+let rec byte8_bounded_strlen_aux (source:list byte8) (count,limit:word16) on source ≝
+ match source with
+  [ nil ⇒ True
+  | cons _ tl ⇒ match eq_w16 count 〈〈xF,xF〉:〈xF,xF〉〉 with
+   [ true ⇒ False
+   | false ⇒ match le_w16 (succ_w16 count) limit with
+    [ true ⇒ byte8_bounded_strlen_aux tl (succ_w16 count) limit
+    | false ⇒ False
+    ]
+   ]
+  ].
+
+definition byte8_bounded_strlen ≝
+λsource:list byte8.λlimit:word16.
+ byte8_bounded_strlen_aux source 〈〈x0,x0〉:〈x0,x0〉〉 limit.
+
+(* strlen esadecimale normale *)
+let rec byte8_strlen_aux (source:list byte8) (count:word16) on source ≝
+ match source with
+  [ nil ⇒ count
+  | cons _ tl ⇒ byte8_strlen_aux tl (succ_w16 count)
+  ].
+
+definition byte8_strlen ≝
+λsource:list byte8.
+ byte8_strlen_aux source 〈〈x0,x0〉:〈x0,x0〉〉.
+
+(* reverse *)
+let rec byte8_reverse (source:list byte8) on source ≝
+ match source return λ_.list byte8 with
+  [ nil ⇒ nil byte8
+  | cons hd tl ⇒ (byte8_reverse tl)@[ hd ]
+  ].
+
+(* hex dump: memory -> list byte8 *)
+let rec byte8_hexdump_aux (t:memory_impl) (mem:aux_mem_type t) (inf:word16) (count:nat) (out:list byte8) on count ≝
+ match count with
+  [ O ⇒ out
+  | S n ⇒ byte8_hexdump_aux t mem (succ_w16 inf) n (out@[ mem_read_abs t mem inf ])
+  ].
+
+definition byte8_hexdump ≝
+λt:memory_impl.λmem:aux_mem_type t.λinf:word16.λcount:nat.
+ byte8_hexdump_aux t mem inf count [].
+
+(* ************* *)
+(* TEST PATTERNS *)
+(* ************* *)
+
+(* lista di 3072 numeri random generati da Mathematica, in blocchi da 256, per velocita' di parsing *)
+definition dTest_random_ex01a ≝
+[
+〈x8,x1〉;〈x5,xE〉;〈x7,x6〉;〈xD,x1〉;〈x7,x5〉;〈x1,x0〉;〈x8,x5〉;〈x9,x0〉;〈x1,xF〉;〈x1,x2〉;〈xE,x2〉;〈x3,xC〉;〈x1,xD〉;〈x0,x6〉;〈x3,xC〉;〈xD,x1〉;
+〈x8,x3〉;〈xE,xB〉;〈x7,x2〉;〈x1,xF〉;〈x9,x0〉;〈xF,x0〉;〈x4,x7〉;〈xA,x3〉;〈xD,x0〉;〈x4,x6〉;〈xC,x5〉;〈x3,x2〉;〈xC,x9〉;〈x0,xB〉;〈x1,xB〉;〈x7,x3〉
+].
+
+definition dTest_random_ex01b ≝
+[
+〈x1,x1〉;〈xA,xF〉;〈xD,x2〉;〈xB,xF〉;〈x0,xB〉;〈xB,xC〉;〈x0,x8〉;〈xE,x5〉;〈xF,xE〉;〈xA,xC〉;〈xF,xC〉;〈x9,x4〉;〈x3,x4〉;〈x8,x8〉;〈x2,x8〉;〈xD,xB〉;
+〈xD,x3〉;〈x3,xD〉;〈x3,x8〉;〈x6,xA〉;〈x4,xD〉;〈x7,x4〉;〈x4,x9〉;〈x7,xE〉;〈x6,xF〉;〈x4,xD〉;〈x3,xD〉;〈x3,xA〉;〈xC,x2〉;〈xD,x3〉;〈x3,x6〉;〈x7,x0〉
+].
+
+definition dTest_random_ex01c ≝
+[
+〈x8,x0〉;〈xC,x9〉;〈x3,xB〉;〈x5,x2〉;〈x8,xF〉;〈x1,xE〉;〈x8,x4〉;〈x5,x2〉;〈x2,xD〉;〈xA,xF〉;〈x1,xB〉;〈x0,x1〉;〈x3,x5〉;〈x7,x2〉;〈x1,x0〉;〈x0,x7〉;
+〈x5,xD〉;〈xA,xB〉;〈xE,x0〉;〈x7,x6〉;〈xB,xD〉;〈x3,xC〉;〈x0,xD〉;〈xC,xE〉;〈xB,x7〉;〈x9,xD〉;〈xA,x9〉;〈xF,x3〉;〈xC,x1〉;〈x9,x3〉;〈x8,xC〉;〈x4,xE〉;
+〈xF,x8〉;〈xC,x3〉;〈x1,x3〉;〈x3,xA〉;〈xA,x2〉;〈xB,xD〉;〈x1,x0〉;〈x9,xB〉;〈x5,xC〉;〈xC,x8〉;〈xE,x5〉;〈xC,x8〉;〈xA,x6〉;〈xF,xB〉;〈x1,x3〉;〈xC,x8〉;
+〈x6,x4〉;〈x0,x9〉;〈xD,x6〉;〈xD,x7〉;〈x3,x9〉;〈xF,x9〉;〈xE,x4〉;〈x0,x3〉;〈x4,x9〉;〈xC,xE〉;〈x7,x1〉;〈x5,x7〉;〈x3,x4〉;〈x2,xE〉;〈xA,x3〉;〈x5,x3〉
+].
+
+definition dTest_random_ex01d ≝
+[
+〈xF,x9〉;〈x2,xA〉;〈xF,x0〉;〈xE,x1〉;〈x6,x6〉;〈x8,x9〉;〈x2,xE〉;〈x8,x6〉;〈xC,x7〉;〈x2,xD〉;〈x1,x2〉;〈xD,x5〉;〈x4,xA〉;〈x7,x9〉;〈x6,xD〉;〈x1,xB〉;
+〈x4,x1〉;〈x9,x4〉;〈x6,xC〉;〈x6,xD〉;〈xF,x4〉;〈x5,xD〉;〈x8,xB〉;〈xB,xE〉;〈xC,x8〉;〈xE,x5〉;〈xA,x4〉;〈xB,xA〉;〈xB,x1〉;〈x2,x1〉;〈x5,x0〉;〈x6,xB〉;
+〈x9,x5〉;〈x0,x6〉;〈x4,x5〉;〈xE,x0〉;〈x0,xF〉;〈x5,xA〉;〈xC,xE〉;〈xC,x8〉;〈x8,x2〉;〈x4,xF〉;〈xC,x2〉;〈xF,x5〉;〈xF,x8〉;〈x1,x3〉;〈xD,xA〉;〈x2,xB〉;
+〈x7,x9〉;〈xB,xF〉;〈xA,x4〉;〈x5,xA〉;〈xD,x2〉;〈x7,x6〉;〈x8,xD〉;〈x1,xF〉;〈x1,x5〉;〈x5,xA〉;〈xD,xE〉;〈x2,x3〉;〈x9,xD〉;〈x7,xE〉;〈x6,x7〉;〈x6,x3〉;
+〈x1,x1〉;〈xF,x5〉;〈x2,x4〉;〈xD,x6〉;〈x6,xA〉;〈x0,xC〉;〈x7,xF〉;〈x1,x2〉;〈xD,x6〉;〈xE,xE〉;〈xB,xA〉;〈x4,x4〉;〈x3,x9〉;〈x9,x2〉;〈x6,x6〉;〈xE,xA〉;
+〈x2,x3〉;〈x4,x4〉;〈xC,xF〉;〈x7,x4〉;〈x6,xB〉;〈x8,x3〉;〈x7,x7〉;〈x0,x4〉;〈x6,x4〉;〈xB,x3〉;〈x3,xC〉;〈x2,x6〉;〈xF,xC〉;〈xD,x5〉;〈x4,x1〉;〈xB,x7〉;
+〈x2,x4〉;〈xE,x0〉;〈xB,x4〉;〈x1,x3〉;〈x3,x2〉;〈x4,x5〉;〈x1,x8〉;〈xD,xB〉;〈x0,x0〉;〈xB,x6〉;〈x5,xF〉;〈x3,xA〉;〈xB,x7〉;〈x4,xF〉;〈xB,x4〉;〈xD,xF〉;
+〈xF,x4〉;〈x1,xA〉;〈x1,x0〉;〈xE,x9〉;〈xC,x5〉;〈x2,xC〉;〈xB,xC〉;〈x5,xB〉;〈x0,x5〉;〈x4,xA〉;〈x5,x2〉;〈xC,x0〉;〈x1,xF〉;〈x5,x9〉;〈xF,x2〉;〈xE,x5〉
+].
+
+definition dTest_random_ex02 ≝
+[
+〈x9,x7〉;〈x5,x1〉;〈xF,x7〉;〈x3,xC〉;〈x7,x7〉;〈x8,x1〉;〈x7,xC〉;〈xB,x0〉;〈x1,xD〉;〈x1,xA〉;〈x7,xB〉;〈x4,x7〉;〈x7,x9〉;〈xC,x2〉;〈x3,xF〉;〈xD,x3〉;
+〈xB,x6〉;〈xD,x7〉;〈xC,x9〉;〈x9,x2〉;〈x0,xD〉;〈x7,xF〉;〈x0,xB〉;〈x7,x5〉;〈x0,xE〉;〈x5,x1〉;〈x9,xF〉;〈x4,xC〉;〈xE,x5〉;〈xD,x0〉;〈xC,x1〉;〈x3,x9〉;
+〈x7,xF〉;〈xE,xA〉;〈x4,x1〉;〈x0,x5〉;〈x0,x8〉;〈x7,xE〉;〈x7,x6〉;〈x2,xF〉;〈x6,x2〉;〈x7,x1〉;〈xF,x8〉;〈x3,x8〉;〈xB,x6〉;〈x7,x0〉;〈xD,xD〉;〈xD,x0〉;
+〈xA,xD〉;〈x7,x8〉;〈x3,xD〉;〈x7,xB〉;〈xC,xC〉;〈x6,xD〉;〈x3,xF〉;〈xA,xC〉;〈xC,x7〉;〈x6,xE〉;〈xF,xF〉;〈x4,x5〉;〈xC,x1〉;〈x7,x4〉;〈xB,xF〉;〈x2,x9〉;
+〈x4,x9〉;〈xC,x5〉;〈x7,xC〉;〈x2,xD〉;〈xF,x2〉;〈xC,xC〉;〈xF,xA〉;〈x5,x8〉;〈xA,xC〉;〈x5,x8〉;〈x5,x1〉;〈x0,xE〉;〈x4,x8〉;〈x7,x0〉;〈x4,x1〉;〈xE,xD〉;
+〈x9,x5〉;〈xB,x4〉;〈x4,x7〉;〈xF,x1〉;〈x8,x1〉;〈xE,x4〉;〈x4,x0〉;〈x5,x4〉;〈x7,xB〉;〈xA,x1〉;〈xD,x2〉;〈x0,x2〉;〈xC,x5〉;〈x7,xD〉;〈x7,xA〉;〈xF,x1〉;
+〈x0,x0〉;〈x8,xD〉;〈x6,x6〉;〈xB,x0〉;〈xA,x9〉;〈x7,xA〉;〈x2,xB〉;〈x2,xC〉;〈x6,xF〉;〈x5,xF〉;〈x6,xF〉;〈x4,xE〉;〈x2,xB〉;〈x1,x1〉;〈xF,xF〉;〈x4,x7〉;
+〈x7,xF〉;〈xC,xF〉;〈xD,x3〉;〈x8,x5〉;〈xB,x8〉;〈x4,xD〉;〈x9,x9〉;〈x8,x3〉;〈x7,x6〉;〈xE,x8〉;〈x3,x0〉;〈xF,x8〉;〈xD,x0〉;〈x5,x9〉;〈xB,x8〉;〈x7,x1〉;
+〈x5,x7〉;〈xA,xF〉;〈xB,x5〉;〈xD,x5〉;〈xF,xC〉;〈x8,x2〉;〈x1,x6〉;〈x1,x0〉;〈xB,x3〉;〈xC,x1〉;〈x1,xA〉;〈x0,x1〉;〈x1,x1〉;〈xF,xF〉;〈x5,x9〉;〈x2,x4〉;
+〈xC,x5〉;〈x7,x2〉;〈xD,x8〉;〈xD,x0〉;〈xA,xD〉;〈xF,xE〉;〈xD,x7〉;〈x1,x1〉;〈xC,xE〉;〈xF,x9〉;〈xB,x8〉;〈x7,x7〉;〈x3,xA〉;〈x1,xF〉;〈x6,x1〉;〈x1,xB〉;
+〈xB,xB〉;〈x5,x5〉;〈x8,x0〉;〈x7,xC〉;〈x2,x5〉;〈x3,x4〉;〈x8,x9〉;〈xF,x2〉;〈xC,x9〉;〈xD,xF〉;〈x3,x5〉;〈xC,x5〉;〈x1,x2〉;〈xF,x0〉;〈x0,x5〉;〈xD,xE〉;
+〈x2,x6〉;〈x4,x9〉;〈xB,x7〉;〈x3,x9〉;〈x0,x5〉;〈xC,x2〉;〈xD,xB〉;〈xF,xC〉;〈x9,xF〉;〈xA,x9〉;〈x6,x6〉;〈xA,xD〉;〈x4,xA〉;〈x3,xF〉;〈xB,xF〉;〈x6,xD〉;
+〈x8,x7〉;〈x6,xA〉;〈xB,x1〉;〈x3,xE〉;〈xB,x6〉;〈x0,xE〉;〈x7,xA〉;〈x3,xB〉;〈x4,x5〉;〈xE,x9〉;〈xC,xE〉;〈x6,xA〉;〈x6,xA〉;〈x7,x0〉;〈x6,x0〉;〈x6,xA〉;
+〈x2,xC〉;〈xD,x2〉;〈xB,x8〉;〈x3,x6〉;〈x2,x1〉;〈x0,x0〉;〈x5,x4〉;〈x3,x1〉;〈x6,x0〉;〈x1,xB〉;〈x4,xC〉;〈xC,xA〉;〈xB,xE〉;〈x5,xF〉;〈x8,x1〉;〈xB,x7〉;
+〈x9,xB〉;〈x2,x6〉;〈x9,x4〉;〈x2,xB〉;〈x4,x1〉;〈x2,xB〉;〈x9,x8〉;〈x6,x3〉;〈x6,x6〉;〈x6,x5〉;〈x4,x6〉;〈x2,x3〉;〈xE,x5〉;〈x0,x7〉;〈x9,xE〉;〈x1,xC〉;
+〈x3,x8〉;〈x5,xC〉;〈x9,x7〉;〈x6,x3〉;〈x5,x3〉;〈x6,x6〉;〈x0,x8〉;〈x5,xD〉;〈x0,x8〉;〈xD,xB〉;〈x6,xE〉;〈x5,x6〉;〈x7,x0〉;〈x3,x2〉;〈x4,x5〉;〈x0,x2〉
+].
+
+definition dTest_random_ex03 ≝
+[
+〈x6,x3〉;〈x7,x2〉;〈x9,xC〉;〈xD,x9〉;〈x5,x0〉;〈x0,x6〉;〈x5,x9〉;〈x1,x7〉;〈x6,x8〉;〈xD,x2〉;〈xD,x7〉;〈x8,xE〉;〈x6,x9〉;〈x5,xF〉;〈x8,x1〉;〈x8,x4〉;
+〈x8,x7〉;〈xD,xC〉;〈x9,x8〉;〈xE,x5〉;〈xB,x5〉;〈xC,x3〉;〈x2,x5〉;〈x6,xC〉;〈x9,x2〉;〈xD,xD〉;〈x2,xA〉;〈xD,x1〉;〈x1,x4〉;〈x7,xE〉;〈x1,x7〉;〈xB,x2〉;
+〈x9,x8〉;〈x5,x5〉;〈xF,xC〉;〈x3,xD〉;〈x8,xD〉;〈xE,xF〉;〈x8,x1〉;〈xB,x8〉;〈xB,xB〉;〈x5,x1〉;〈x0,x0〉;〈xB,x4〉;〈x2,xE〉;〈x3,x0〉;〈x6,x0〉;〈x7,xE〉;
+〈x9,x0〉;〈xE,x3〉;〈xF,x4〉;〈x7,x2〉;〈x1,xC〉;〈xB,x3〉;〈x7,x8〉;〈x1,xB〉;〈x9,xF〉;〈x1,xB〉;〈x0,x3〉;〈xA,x3〉;〈x0,x5〉;〈xD,xE〉;〈x3,x8〉;〈xB,xA〉;
+〈x0,xE〉;〈xE,xD〉;〈xE,xC〉;〈x1,xF〉;〈x3,x8〉;〈xE,x3〉;〈xF,x7〉;〈xA,xA〉;〈xE,x9〉;〈x3,xD〉;〈xF,xF〉;〈xF,x3〉;〈x1,x4〉;〈x2,xC〉;〈x8,x8〉;〈x6,x1〉;
+〈x3,x0〉;〈xA,xB〉;〈x1,x8〉;〈xD,xC〉;〈xF,xE〉;〈x6,xA〉;〈x2,x9〉;〈xF,x1〉;〈xC,xB〉;〈x9,x0〉;〈x7,x8〉;〈x9,x9〉;〈x1,xF〉;〈x2,x8〉;〈xF,x9〉;〈xC,xB〉;
+〈x1,x4〉;〈x8,x4〉;〈xF,x3〉;〈xD,x6〉;〈x7,xE〉;〈xE,xC〉;〈x5,x6〉;〈xC,xE〉;〈xD,xA〉;〈x5,xE〉;〈x6,x1〉;〈xF,x1〉;〈x6,x6〉;〈x6,x9〉;〈x9,x3〉;〈x5,x9〉;
+〈x3,xC〉;〈x1,xD〉;〈x6,xB〉;〈xF,x4〉;〈x5,x9〉;〈x4,xD〉;〈x3,x8〉;〈xA,x9〉;〈x3,xB〉;〈x7,xF〉;〈xB,x2〉;〈xE,xC〉;〈xA,xE〉;〈xF,x6〉;〈xB,x2〉;〈x2,x2〉;
+〈x6,x4〉;〈x2,x7〉;〈x6,xC〉;〈x2,x0〉;〈xE,xE〉;〈x5,x1〉;〈x3,xE〉;〈x8,x8〉;〈xD,xD〉;〈xC,x1〉;〈xD,xC〉;〈xC,x1〉;〈x6,x6〉;〈x6,x1〉;〈x4,x2〉;〈x7,x7〉;
+〈x3,x6〉;〈x0,x8〉;〈x2,x9〉;〈x6,x0〉;〈xA,x9〉;〈xF,xC〉;〈x7,xC〉;〈xA,x7〉;〈xB,x4〉;〈xF,xC〉;〈x8,x7〉;〈x1,xD〉;〈x6,xC〉;〈xA,x2〉;〈x3,xF〉;〈x1,xD〉;
+〈x1,x7〉;〈x0,xF〉;〈x0,x2〉;〈x2,x6〉;〈xA,x2〉;〈x6,xA〉;〈x5,xC〉;〈xE,xD〉;〈x2,x7〉;〈xC,x5〉;〈x7,xB〉;〈xF,x5〉;〈x9,xC〉;〈x8,x5〉;〈x6,x3〉;〈x5,x6〉;
+〈xC,x3〉;〈x4,xB〉;〈x1,xB〉;〈xA,x0〉;〈x1,xB〉;〈x8,x9〉;〈x3,x5〉;〈xD,x6〉;〈xD,x9〉;〈xD,xD〉;〈x2,xE〉;〈x6,x2〉;〈x7,x5〉;〈xE,x7〉;〈x1,x8〉;〈x4,xD〉;
+〈xD,x7〉;〈x5,x8〉;〈xA,x7〉;〈x5,xF〉;〈x9,x4〉;〈x8,x7〉;〈xA,x8〉;〈xE,x7〉;〈x2,xB〉;〈xF,x2〉;〈xE,x7〉;〈xB,x9〉;〈x0,x6〉;〈xA,xF〉;〈xD,xA〉;〈xD,xC〉;
+〈xC,x6〉;〈x3,xF〉;〈x8,xD〉;〈x7,x9〉;〈x9,x5〉;〈xD,xA〉;〈x5,xB〉;〈x9,x2〉;〈xE,xE〉;〈x3,xC〉;〈xF,xE〉;〈x4,x9〉;〈x5,xA〉;〈x1,x0〉;〈x4,xD〉;〈x8,x9〉;
+〈x8,x3〉;〈x2,x6〉;〈xE,xC〉;〈x8,xD〉;〈xC,x9〉;〈x7,x7〉;〈xE,xE〉;〈xF,x1〉;〈x4,x0〉;〈x6,xD〉;〈x4,x9〉;〈x5,x7〉;〈x9,xB〉;〈xC,x4〉;〈x1,xF〉;〈x8,x0〉;
+〈x9,x5〉;〈xB,xC〉;〈xE,x8〉;〈xF,x9〉;〈xD,x7〉;〈x1,x4〉;〈x3,xE〉;〈xC,x3〉;〈x6,xF〉;〈x8,xF〉;〈x7,x2〉;〈xD,x5〉;〈xB,xE〉;〈x8,xA〉;〈xA,x3〉;〈xF,x7〉
+].
+
+definition dTest_random_ex04 ≝
+[
+〈x6,x0〉;〈x3,xA〉;〈x7,x4〉;〈xF,xB〉;〈xB,xD〉;〈x7,x4〉;〈x8,x3〉;〈xE,x3〉;〈x9,xD〉;〈xD,x9〉;〈xB,x8〉;〈x1,x3〉;〈x5,x0〉;〈x4,x0〉;〈x8,xA〉;〈x9,x6〉;
+〈x3,xA〉;〈xA,x6〉;〈xE,xC〉;〈x7,xC〉;〈x1,x5〉;〈x8,x7〉;〈x4,xD〉;〈x6,xA〉;〈xA,xA〉;〈xE,x0〉;〈xB,xA〉;〈xF,xF〉;〈x3,xB〉;〈xE,x2〉;〈x5,x1〉;〈x2,x2〉;
+〈x2,x2〉;〈x1,xF〉;〈xA,x1〉;〈x2,x1〉;〈xA,xF〉;〈x3,x7〉;〈x8,xA〉;〈xD,xF〉;〈xE,x3〉;〈x6,x9〉;〈xE,xE〉;〈xC,x4〉;〈xE,x7〉;〈x7,x1〉;〈x9,x6〉;〈x1,x1〉;
+〈xE,x4〉;〈x3,x9〉;〈xE,x5〉;〈xA,xF〉;〈xF,x5〉;〈x5,x7〉;〈xE,xB〉;〈x5,x5〉;〈x6,x5〉;〈x8,xB〉;〈x3,xE〉;〈x8,xD〉;〈x4,x6〉;〈x5,x3〉;〈xB,x2〉;〈x1,x9〉;
+〈x3,x4〉;〈xE,x9〉;〈x4,xA〉;〈x4,xB〉;〈x5,x2〉;〈x3,x0〉;〈x3,xF〉;〈xA,x7〉;〈x4,xF〉;〈x1,xA〉;〈xB,x8〉;〈x6,x4〉;〈x5,xB〉;〈xD,x9〉;〈x6,xD〉;〈x6,x1〉;
+〈xA,x5〉;〈xC,xF〉;〈x8,xC〉;〈xD,xD〉;〈xE,x6〉;〈xD,x5〉;〈x3,x6〉;〈x0,xC〉;〈x8,xD〉;〈xF,x7〉;〈x4,xE〉;〈x9,xC〉;〈xB,xF〉;〈x2,xB〉;〈x4,x4〉;〈xD,x1〉;
+〈xC,x0〉;〈x8,x0〉;〈x0,x8〉;〈xA,xD〉;〈xC,xE〉;〈xB,xD〉;〈x4,xC〉;〈x5,x3〉;〈x6,x5〉;〈xB,x6〉;〈x4,x8〉;〈xF,x6〉;〈x6,x4〉;〈x7,xC〉;〈x9,x8〉;〈x1,x0〉;
+〈x9,xD〉;〈xF,xD〉;〈x4,x9〉;〈xC,x4〉;〈xD,xD〉;〈x1,x4〉;〈xB,x6〉;〈x6,xF〉;〈x3,xB〉;〈x4,x6〉;〈xD,x7〉;〈x1,xA〉;〈x4,x4〉;〈xA,x4〉;〈x8,x1〉;〈x3,x1〉;
+〈xA,x2〉;〈x4,x0〉;〈x7,x0〉;〈x3,x9〉;〈x9,xA〉;〈x4,xC〉;〈x4,xF〉;〈x9,x3〉;〈x9,xD〉;〈xD,x4〉;〈x9,x7〉;〈x3,x9〉;〈xA,x8〉;〈xA,x8〉;〈xF,x9〉;〈xB,x3〉;
+〈xE,x7〉;〈xD,x8〉;〈x4,xD〉;〈x6,xD〉;〈x8,x8〉;〈x6,xB〉;〈x5,x4〉;〈x5,x5〉;〈x9,x2〉;〈x1,x9〉;〈xE,x4〉;〈xB,x1〉;〈xE,x3〉;〈x4,x6〉;〈xD,x5〉;〈x6,xC〉;
+〈xF,x9〉;〈x3,x3〉;〈xE,x5〉;〈xD,x2〉;〈x7,x7〉;〈x1,x8〉;〈x8,x6〉;〈x1,x4〉;〈x7,xE〉;〈x1,xE〉;〈xC,x8〉;〈xB,x4〉;〈xC,xE〉;〈xC,x7〉;〈x5,x7〉;〈x2,xD〉;
+〈x5,x2〉;〈xE,x7〉;〈x5,x7〉;〈x9,xB〉;〈x1,xA〉;〈x4,x9〉;〈x0,x9〉;〈x4,xB〉;〈xF,x6〉;〈xB,x5〉;〈x0,xA〉;〈x2,xC〉;〈xB,xA〉;〈x1,xF〉;〈x2,x6〉;〈x3,x8〉;
+〈x6,x0〉;〈x3,x4〉;〈xC,x4〉;〈x8,x9〉;〈xB,xF〉;〈x0,x4〉;〈x5,xF〉;〈xF,x4〉;〈x1,x1〉;〈xF,x3〉;〈x4,xA〉;〈xE,x2〉;〈xD,x8〉;〈x3,x6〉;〈xF,xA〉;〈x4,x3〉;
+〈x0,xC〉;〈x6,x0〉;〈x5,x1〉;〈x2,xD〉;〈x0,x3〉;〈xF,xC〉;〈x9,x2〉;〈x2,x1〉;〈xC,xA〉;〈x7,x5〉;〈x6,x6〉;〈xD,xC〉;〈x8,x4〉;〈xC,x4〉;〈x6,xF〉;〈xF,x0〉;
+〈x0,x4〉;〈x5,x2〉;〈xA,xB〉;〈x9,xA〉;〈xC,xE〉;〈xA,xA〉;〈x2,xF〉;〈x3,x6〉;〈xF,x3〉;〈xA,x7〉;〈x4,x2〉;〈x0,x7〉;〈x5,xE〉;〈xB,x6〉;〈x5,xB〉;〈x9,x8〉;
+〈x7,x9〉;〈x4,x3〉;〈x8,x8〉;〈x9,xA〉;〈x0,xA〉;〈x5,x5〉;〈xE,x3〉;〈x9,x8〉;〈x7,x5〉;〈xA,x2〉;〈xE,xA〉;〈xB,x9〉;〈x3,x2〉;〈x4,x1〉;〈x0,x7〉;〈x3,x8〉
+].
+
+definition dTest_random_ex05 ≝
+[
+〈xB,x1〉;〈x0,x3〉;〈x4,xE〉;〈x6,xF〉;〈x9,x7〉;〈x3,xC〉;〈x5,x4〉;〈xB,xB〉;〈xB,xC〉;〈x0,xB〉;〈x0,xC〉;〈xA,xB〉;〈x3,xB〉;〈x2,x8〉;〈xA,xA〉;〈x0,x4〉;
+〈x5,x0〉;〈xA,x3〉;〈x6,xB〉;〈xF,xA〉;〈xF,x0〉;〈xE,x6〉;〈x1,x8〉;〈xD,x8〉;〈x7,xA〉;〈x5,xF〉;〈x0,x3〉;〈x8,x2〉;〈x9,x6〉;〈x9,xA〉;〈xB,xF〉;〈x1,x5〉;
+〈x8,x5〉;〈xE,xF〉;〈xB,x3〉;〈x7,xB〉;〈xE,xE〉;〈xE,xF〉;〈x1,x0〉;〈x6,x9〉;〈xC,x2〉;〈xF,x9〉;〈x2,xD〉;〈x8,x1〉;〈xF,x3〉;〈x0,xE〉;〈xC,x3〉;〈x7,xF〉;
+〈xD,xD〉;〈x2,xE〉;〈x2,xA〉;〈xB,xD〉;〈xE,xA〉;〈x9,x2〉;〈xF,xF〉;〈xF,xA〉;〈x7,xB〉;〈xD,x3〉;〈x3,x0〉;〈x7,x5〉;〈x6,x7〉;〈xA,x8〉;〈x0,xF〉;〈x2,x1〉;
+〈xD,xC〉;〈xE,x5〉;〈xE,x2〉;〈x8,x9〉;〈x2,x9〉;〈xC,x5〉;〈xA,x3〉;〈xA,x2〉;〈x4,x2〉;〈x3,xF〉;〈xA,x3〉;〈x5,x8〉;〈xE,x0〉;〈x7,xC〉;〈x0,x3〉;〈xF,xF〉;
+〈x2,x8〉;〈x8,xB〉;〈x8,xB〉;〈x1,x2〉;〈xD,x8〉;〈xA,x8〉;〈x7,x6〉;〈xB,x9〉;〈xE,x2〉;〈xF,xE〉;〈x2,x1〉;〈x3,xF〉;〈xA,xC〉;〈x4,x6〉;〈xB,xC〉;〈xF,x8〉;
+〈xD,x3〉;〈xE,xB〉;〈xF,xC〉;〈x9,xF〉;〈xE,x7〉;〈x6,x1〉;〈xC,xB〉;〈xB,xF〉;〈x4,xE〉;〈xC,x4〉;〈x9,x7〉;〈x1,xE〉;〈x0,xD〉;〈x7,x9〉;〈x8,x3〉;〈xA,xB〉;
+〈x4,xC〉;〈x2,x6〉;〈x6,x3〉;〈x6,xF〉;〈xE,xE〉;〈x5,x9〉;〈x8,x1〉;〈x0,x2〉;〈x2,xC〉;〈xE,xD〉;〈x6,xF〉;〈x0,x4〉;〈x1,x0〉;〈xE,x0〉;〈xD,xA〉;〈xB,xE〉;
+〈xE,xE〉;〈x5,x7〉;〈xB,x0〉;〈x3,x1〉;〈x4,x1〉;〈xD,xC〉;〈x3,xC〉;〈xC,xC〉;〈x5,x8〉;〈x2,x8〉;〈x2,xC〉;〈x1,xB〉;〈x8,x6〉;〈xD,x6〉;〈xF,x9〉;〈xD,x5〉;
+〈x4,xA〉;〈xE,xA〉;〈x0,xB〉;〈x2,x0〉;〈x2,xC〉;〈x4,x2〉;〈xC,xE〉;〈x4,x5〉;〈x2,xB〉;〈x0,x1〉;〈xA,xA〉;〈xB,x1〉;〈x6,xE〉;〈xB,x7〉;〈xB,x7〉;〈x2,x8〉;
+〈x9,x5〉;〈x1,x9〉;〈xA,x7〉;〈x5,xC〉;〈x4,xE〉;〈x7,xB〉;〈x3,xE〉;〈xD,x3〉;〈x9,x0〉;〈x8,x6〉;〈x7,x1〉;〈x1,x4〉;〈xD,x2〉;〈xD,x4〉;〈xB,x4〉;〈xF,x2〉;
+〈x3,x1〉;〈x2,x8〉;〈x4,x5〉;〈xF,xD〉;〈x7,x8〉;〈x5,xD〉;〈xF,xA〉;〈xF,x3〉;〈x9,x5〉;〈x4,xD〉;〈x3,x1〉;〈xB,x8〉;〈xC,xC〉;〈x2,x1〉;〈x1,x9〉;〈x4,x2〉;
+〈x2,xA〉;〈xF,x2〉;〈xB,xA〉;〈x0,x4〉;〈x9,xF〉;〈x4,x3〉;〈x4,x5〉;〈x1,xC〉;〈x7,x4〉;〈xB,xB〉;〈x7,x0〉;〈x5,xE〉;〈x0,x1〉;〈xA,xC〉;〈x6,xD〉;〈xD,x7〉;
+〈x9,xC〉;〈x9,xD〉;〈x1,xA〉;〈x9,x8〉;〈xB,x1〉;〈xF,xC〉;〈x6,x1〉;〈xA,x3〉;〈x4,x1〉;〈x4,x1〉;〈xA,xF〉;〈x1,xD〉;〈xE,x1〉;〈x3,x2〉;〈x1,x9〉;〈x6,x0〉;
+〈x2,x9〉;〈x9,x7〉;〈x8,x5〉;〈x5,x3〉;〈x5,x3〉;〈x9,x1〉;〈xB,x3〉;〈x9,x4〉;〈xD,x5〉;〈x9,xD〉;〈x4,xC〉;〈x3,x6〉;〈x0,xE〉;〈x8,x4〉;〈xA,x1〉;〈x4,x6〉;
+〈x6,xA〉;〈x1,xF〉;〈xF,x3〉;〈x6,xB〉;〈xB,xE〉;〈x4,xA〉;〈x1,x9〉;〈x7,x5〉;〈xF,xC〉;〈xC,x6〉;〈xE,xA〉;〈x7,xE〉;〈xD,x1〉;〈x3,x3〉;〈x6,x7〉;〈xB,x7〉
+].
+
+definition dTest_random_ex06 ≝
+[
+〈xE,xE〉;〈x5,x9〉;〈xE,x2〉;〈xD,xD〉;〈x2,x2〉;〈x8,xC〉;〈x9,xB〉;〈x3,xE〉;〈x9,x8〉;〈xF,xC〉;〈x1,x3〉;〈xE,x2〉;〈x0,xC〉;〈x4,xE〉;〈x3,x1〉;〈x8,x7〉;
+〈x6,x7〉;〈x6,xA〉;〈x4,xC〉;〈x4,xC〉;〈x7,x2〉;〈x0,x0〉;〈x0,x5〉;〈x1,xF〉;〈xF,x6〉;〈x3,x0〉;〈xE,xE〉;〈xD,xE〉;〈xB,x1〉;〈x4,xC〉;〈xF,x7〉;〈xE,xC〉;
+〈x2,xC〉;〈x4,x0〉;〈x6,xB〉;〈x6,x8〉;〈x9,x0〉;〈x8,x8〉;〈x6,xF〉;〈xB,x3〉;〈x4,x7〉;〈x6,x2〉;〈x9,x2〉;〈x9,xB〉;〈x2,xB〉;〈x3,x2〉;〈x4,x0〉;〈xA,x7〉;
+〈x8,x9〉;〈x4,x0〉;〈x2,x3〉;〈x5,xC〉;〈xF,x9〉;〈x2,x9〉;〈x6,x2〉;〈xA,xE〉;〈x5,xB〉;〈xC,x9〉;〈x2,xC〉;〈x9,x2〉;〈x6,xF〉;〈xF,x5〉;〈xA,x0〉;〈x0,xE〉;
+〈xD,xE〉;〈xF,x9〉;〈x0,x9〉;〈x1,x0〉;〈x3,x9〉;〈x4,x6〉;〈xC,x5〉;〈xE,x2〉;〈x8,x3〉;〈xD,x5〉;〈x8,xE〉;〈x4,x6〉;〈x4,xC〉;〈xA,xC〉;〈x7,xF〉;〈x4,xF〉;
+〈xC,x1〉;〈x4,xF〉;〈x1,xA〉;〈x6,x1〉;〈x9,x6〉;〈x0,xB〉;〈x0,x0〉;〈x6,xF〉;〈x2,x6〉;〈x8,xC〉;〈xE,xE〉;〈x9,x3〉;〈x1,xB〉;〈x9,xE〉;〈xA,x5〉;〈x9,x6〉;
+〈x2,xA〉;〈xE,xB〉;〈x4,x6〉;〈x5,xF〉;〈x3,xC〉;〈xD,x6〉;〈x2,xD〉;〈x9,x4〉;〈x6,xB〉;〈xF,x4〉;〈xD,xA〉;〈x6,x9〉;〈x5,x9〉;〈xA,xC〉;〈xB,xD〉;〈x9,xE〉;
+〈x4,x8〉;〈x0,x2〉;〈xD,xC〉;〈x5,xC〉;〈x6,x0〉;〈x2,xA〉;〈x6,xE〉;〈xC,xA〉;〈x6,xE〉;〈x1,xF〉;〈xD,x4〉;〈x3,xA〉;〈xB,x0〉;〈x9,xE〉;〈x8,xF〉;〈xA,xB〉;
+〈xB,x2〉;〈x0,x2〉;〈x4,x7〉;〈x7,xD〉;〈xA,xB〉;〈xD,xB〉;〈xB,x5〉;〈x6,xD〉;〈xE,x2〉;〈x8,x9〉;〈x4,xD〉;〈x0,x4〉;〈xB,xE〉;〈xF,xA〉;〈x2,x2〉;〈x1,x4〉;
+〈x7,x1〉;〈x1,x2〉;〈x1,xB〉;〈x0,xD〉;〈xB,xA〉;〈x5,xA〉;〈x6,xC〉;〈x1,xE〉;〈x3,xA〉;〈x0,xF〉;〈x6,xE〉;〈x4,x4〉;〈xC,x8〉;〈xB,x5〉;〈x8,xC〉;〈x0,x3〉;
+〈x0,x6〉;〈x6,x4〉;〈x8,x5〉;〈x2,x8〉;〈x6,x4〉;〈x2,x2〉;〈x8,x1〉;〈x7,x6〉;〈xF,xE〉;〈xF,xA〉;〈x6,x2〉;〈x9,x1〉;〈xB,xE〉;〈xB,xC〉;〈x6,x1〉;〈x4,xB〉;
+〈x7,xE〉;〈x5,x0〉;〈xB,xC〉;〈xE,xE〉;〈x6,x3〉;〈xC,xF〉;〈x1,xD〉;〈xF,xD〉;〈x6,x2〉;〈x5,xC〉;〈x8,x5〉;〈x9,xE〉;〈xA,x5〉;〈x2,x6〉;〈xE,x7〉;〈x4,x6〉;
+〈x3,xB〉;〈xE,xA〉;〈xB,xE〉;〈x0,x4〉;〈x8,x8〉;〈xF,x2〉;〈x9,x2〉;〈x0,xB〉;〈xD,x9〉;〈xE,x9〉;〈x2,x9〉;〈x3,x8〉;〈x8,x8〉;〈x8,xA〉;〈x6,x9〉;〈x1,x7〉;
+〈x4,xB〉;〈xB,xF〉;〈x0,xC〉;〈xF,x2〉;〈xF,xD〉;〈x7,x3〉;〈x5,x9〉;〈xB,xE〉;〈x5,x4〉;〈x1,xC〉;〈xD,x3〉;〈x3,x1〉;〈x6,x2〉;〈x1,xB〉;〈xB,x7〉;〈x3,x2〉;
+〈xA,x4〉;〈xF,x1〉;〈x7,x0〉;〈x9,xA〉;〈x4,x6〉;〈xA,x1〉;〈x1,xC〉;〈x0,x4〉;〈x6,xC〉;〈xF,x2〉;〈xE,x6〉;〈xC,x1〉;〈xA,x4〉;〈xF,x2〉;〈x2,xA〉;〈x4,xB〉;
+〈x3,x5〉;〈x9,xB〉;〈x9,x9〉;〈xF,xF〉;〈x0,x1〉;〈x1,x3〉;〈xF,x9〉;〈x5,xC〉;〈x3,xC〉;〈x5,x1〉;〈x8,xA〉;〈xA,x5〉;〈x5,xF〉;〈x9,xE〉;〈x5,xE〉;〈xC,x6〉
+].
+
+definition dTest_random_ex07 ≝
+[
+〈x2,x1〉;〈x3,x7〉;〈xD,x2〉;〈xB,x9〉;〈x9,x8〉;〈xA,x1〉;〈x6,x0〉;〈xE,x9〉;〈x4,x5〉;〈xC,xA〉;〈xD,x7〉;〈xB,xD〉;〈xC,xF〉;〈x0,xF〉;〈x2,x4〉;〈xE,x5〉;
+〈x7,x9〉;〈x4,xB〉;〈x1,xC〉;〈x5,x7〉;〈x3,xA〉;〈x2,x4〉;〈x2,x2〉;〈x0,x8〉;〈x3,x3〉;〈xE,x2〉;〈xA,x2〉;〈x5,x8〉;〈x2,x5〉;〈x5,x4〉;〈x7,x1〉;〈x2,xB〉;
+〈xF,xF〉;〈xE,xD〉;〈x4,x8〉;〈xF,x6〉;〈x2,x3〉;〈x3,x1〉;〈xB,xA〉;〈x5,x1〉;〈x9,xF〉;〈xA,xA〉;〈xC,xC〉;〈x0,x3〉;〈x1,x5〉;〈xC,x7〉;〈x2,xD〉;〈xD,x3〉;
+〈xE,xB〉;〈x8,xF〉;〈x8,x4〉;〈x4,x0〉;〈x5,x3〉;〈xA,xD〉;〈x6,x7〉;〈xE,xC〉;〈xA,xF〉;〈xD,xC〉;〈x1,xC〉;〈x7,x4〉;〈x6,xB〉;〈xA,xD〉;〈xC,xD〉;〈xA,x7〉;
+〈x1,x1〉;〈x1,x0〉;〈xC,xF〉;〈xB,xE〉;〈xA,x1〉;〈x0,x1〉;〈x3,xF〉;〈xC,x0〉;〈x8,x5〉;〈x2,x8〉;〈x6,xB〉;〈xC,x3〉;〈x6,xD〉;〈xD,x8〉;〈x7,x5〉;〈x5,xA〉;
+〈xF,x0〉;〈x2,x2〉;〈x4,xB〉;〈x9,xC〉;〈x3,x1〉;〈xE,x4〉;〈xE,x7〉;〈xC,x6〉;〈xF,xC〉;〈x3,x0〉;〈xD,x5〉;〈xF,x9〉;〈x1,xA〉;〈x4,x0〉;〈x1,xF〉;〈x6,xD〉;
+〈xD,x5〉;〈x7,x8〉;〈xB,x5〉;〈x7,x6〉;〈xC,x9〉;〈xE,x1〉;〈xD,xF〉;〈x1,x2〉;〈x6,x1〉;〈xD,xF〉;〈x9,xF〉;〈x5,x7〉;〈x7,xD〉;〈x0,xB〉;〈xA,xD〉;〈x5,xA〉;
+〈xA,x1〉;〈x8,x4〉;〈xE,x5〉;〈xF,x7〉;〈xB,xC〉;〈xD,x3〉;〈xA,x5〉;〈xB,x4〉;〈x8,x5〉;〈x6,x7〉;〈x3,x6〉;〈xF,xC〉;〈xB,x1〉;〈xB,x3〉;〈xC,xB〉;〈x1,xE〉;
+〈xE,xC〉;〈x6,xE〉;〈xE,x1〉;〈x1,xC〉;〈xA,x5〉;〈x5,x3〉;〈x9,x8〉;〈xF,x6〉;〈xD,xF〉;〈x4,x1〉;〈x1,x3〉;〈x2,xE〉;〈x7,xF〉;〈x0,xE〉;〈x3,x8〉;〈x3,xC〉;
+〈xD,x4〉;〈x8,xC〉;〈x2,xA〉;〈x2,x8〉;〈x4,xE〉;〈x7,xE〉;〈x0,xE〉;〈xF,x7〉;〈xC,xA〉;〈x3,xE〉;〈xE,x4〉;〈xB,x4〉;〈x0,x5〉;〈x5,x8〉;〈xD,xC〉;〈x7,x8〉;
+〈xD,x9〉;〈xF,x9〉;〈x7,x9〉;〈x8,x4〉;〈x0,x2〉;〈x3,xF〉;〈xC,xF〉;〈x3,x8〉;〈xD,x7〉;〈x2,x6〉;〈x1,xD〉;〈x1,x8〉;〈x4,xD〉;〈xE,xA〉;〈x7,xA〉;〈xD,x4〉;
+〈x2,x4〉;〈x0,xD〉;〈x4,xD〉;〈x9,x0〉;〈x1,x7〉;〈x1,xE〉;〈x6,xE〉;〈xB,x6〉;〈xC,xC〉;〈xC,x0〉;〈xB,x0〉;〈x5,xE〉;〈x9,x9〉;〈x6,xD〉;〈xC,xF〉;〈xE,xE〉;
+〈x2,x9〉;〈xC,xF〉;〈xA,x2〉;〈x0,xC〉;〈xA,xB〉;〈x7,x4〉;〈x2,x9〉;〈x4,xE〉;〈x8,x2〉;〈x9,x3〉;〈x6,x9〉;〈x7,xB〉;〈xE,xC〉;〈xC,x7〉;〈x8,x9〉;〈xC,xA〉;
+〈xD,xD〉;〈xA,xC〉;〈x6,x5〉;〈x8,x0〉;〈x1,x4〉;〈x0,x9〉;〈x5,x5〉;〈x0,xE〉;〈x8,x4〉;〈x5,xE〉;〈x6,xF〉;〈x3,x4〉;〈x1,x8〉;〈xC,x9〉;〈x8,xB〉;〈xE,x4〉;
+〈x3,xE〉;〈xC,x4〉;〈x1,x6〉;〈xA,x4〉;〈x1,x9〉;〈x3,x4〉;〈x0,xE〉;〈x5,xE〉;〈xF,x9〉;〈x0,x3〉;〈x1,x3〉;〈x7,x2〉;〈x2,x7〉;〈x2,x8〉;〈xA,x7〉;〈x6,xD〉;
+〈xC,x1〉;〈x1,xD〉;〈xF,x0〉;〈x2,x8〉;〈xF,xB〉;〈xF,x6〉;〈x3,x8〉;〈x0,x1〉;〈xF,x9〉;〈xB,xC〉;〈x6,x6〉;〈xF,x8〉;〈x6,xE〉;〈xD,x1〉;〈xB,x5〉;〈x3,x8〉
+].
+
+definition dTest_random_ex08 ≝
+[
+〈x4,x3〉;〈xB,x6〉;〈x6,x8〉;〈xA,xC〉;〈x0,x9〉;〈xF,xD〉;〈x0,x9〉;〈x6,x8〉;〈xE,x0〉;〈x2,x2〉;〈xA,xF〉;〈x4,x0〉;〈x2,x6〉;〈x0,xC〉;〈x5,x2〉;〈xA,x7〉;
+〈xA,xD〉;〈xC,x3〉;〈x8,x2〉;〈xD,xC〉;〈x3,xC〉;〈x6,x5〉;〈xF,x2〉;〈xE,x8〉;〈xC,x0〉;〈x0,x6〉;〈x6,x4〉;〈xB,x1〉;〈x2,x0〉;〈x9,x5〉;〈x2,x2〉;〈xD,xD〉;
+〈xA,xD〉;〈xF,xF〉;〈x1,xB〉;〈x8,xB〉;〈xB,x6〉;〈x4,xA〉;〈xB,xB〉;〈x9,x8〉;〈x1,xA〉;〈xE,xC〉;〈x7,xB〉;〈xA,x6〉;〈x2,xC〉;〈xE,x1〉;〈xC,x7〉;〈xD,xC〉;
+〈x1,x9〉;〈x0,x6〉;〈x0,xA〉;〈x9,xF〉;〈x5,x2〉;〈x2,xB〉;〈xC,xA〉;〈x2,xF〉;〈x4,x0〉;〈xF,x8〉;〈xE,xA〉;〈x8,x7〉;〈x8,x9〉;〈xF,xD〉;〈x5,xD〉;〈x0,x0〉;
+〈x6,xE〉;〈x0,x0〉;〈x0,xD〉;〈x3,x0〉;〈x4,x3〉;〈x5,xA〉;〈x8,xF〉;〈x8,xA〉;〈xA,x4〉;〈x5,x0〉;〈x8,xF〉;〈x0,xC〉;〈x7,x7〉;〈xF,x2〉;〈x6,x5〉;〈xE,x4〉;
+〈x2,xD〉;〈xE,x5〉;〈xA,x8〉;〈x7,xF〉;〈x7,x8〉;〈xE,x3〉;〈x9,x5〉;〈xD,xA〉;〈x0,x7〉;〈x2,x9〉;〈x5,x1〉;〈x9,x4〉;〈xE,x4〉;〈x0,x1〉;〈xB,xF〉;〈x6,xE〉;
+〈x9,x8〉;〈x9,xC〉;〈x9,x0〉;〈xA,x8〉;〈x0,xA〉;〈x3,xD〉;〈x3,xC〉;〈x5,x0〉;〈xE,xB〉;〈x1,x2〉;〈xC,x4〉;〈x5,xF〉;〈x4,x7〉;〈x7,xB〉;〈x2,xC〉;〈xD,xF〉;
+〈x7,x8〉;〈x1,x3〉;〈x7,x4〉;〈xE,x0〉;〈x7,xB〉;〈x7,x1〉;〈x4,x7〉;〈x4,x8〉;〈x1,xB〉;〈xE,x3〉;〈x6,xB〉;〈x0,xB〉;〈x4,xB〉;〈x5,x9〉;〈x9,x3〉;〈xD,xF〉;
+〈xE,x1〉;〈x1,xB〉;〈xD,x0〉;〈xE,xD〉;〈x4,x7〉;〈x4,xD〉;〈xC,x2〉;〈xD,xE〉;〈x5,xC〉;〈xD,xA〉;〈x9,x5〉;〈xC,x8〉;〈x1,x0〉;〈x7,x7〉;〈x7,xF〉;〈xC,x0〉;
+〈xA,x7〉;〈xD,x3〉;〈xD,x3〉;〈xD,x8〉;〈x3,x4〉;〈xA,x1〉;〈x1,x5〉;〈xE,x0〉;〈x0,x4〉;〈x1,xE〉;〈x8,x2〉;〈xC,xA〉;〈xD,x9〉;〈x1,x1〉;〈xB,x1〉;〈xC,x9〉;
+〈x4,xC〉;〈x4,xB〉;〈x0,x9〉;〈x4,x8〉;〈xF,xC〉;〈xD,xD〉;〈x6,xE〉;〈xC,xA〉;〈x7,x6〉;〈xA,xE〉;〈x8,xE〉;〈x3,xB〉;〈xF,xB〉;〈x6,x5〉;〈x8,x3〉;〈x1,xD〉;
+〈xD,xB〉;〈xA,xE〉;〈x4,xF〉;〈xC,x6〉;〈x1,xE〉;〈xC,x5〉;〈xC,xC〉;〈x7,xC〉;〈x2,x8〉;〈xF,x9〉;〈xD,x2〉;〈x8,x6〉;〈x1,x5〉;〈xF,xA〉;〈x4,x1〉;〈x4,x5〉;
+〈x2,xE〉;〈x9,x5〉;〈xB,xF〉;〈x0,xD〉;〈x8,xB〉;〈x8,xD〉;〈x1,x1〉;〈x9,xC〉;〈xB,x8〉;〈xF,xB〉;〈x2,x6〉;〈xD,x6〉;〈x9,x1〉;〈x0,xD〉;〈xC,xD〉;〈x0,x7〉;
+〈x5,x0〉;〈xF,xA〉;〈x2,x9〉;〈x3,xF〉;〈x0,xC〉;〈x2,xB〉;〈xF,xE〉;〈x9,x7〉;〈x5,x5〉;〈x5,xA〉;〈x6,xD〉;〈x9,x6〉;〈x0,x5〉;〈x0,x9〉;〈x4,x5〉;〈xE,xF〉;
+〈x0,xF〉;〈x7,x4〉;〈x9,x3〉;〈x6,xC〉;〈x8,x2〉;〈x3,x7〉;〈xE,xB〉;〈x5,x0〉;〈xF,x5〉;〈xC,x4〉;〈x0,xB〉;〈x3,x8〉;〈x2,xD〉;〈x8,xA〉;〈x9,x3〉;〈x6,xD〉;
+〈x1,xD〉;〈xE,x5〉;〈xF,x7〉;〈xE,x7〉;〈xD,x7〉;〈x5,xC〉;〈xB,x4〉;〈x5,x0〉;〈x7,x5〉;〈x0,xD〉;〈xF,x3〉;〈xC,xE〉;〈x3,x1〉;〈xF,x1〉;〈x8,xE〉;〈x8,xF〉
+].
+
+definition dTest_random_ex09 ≝
+[
+〈xD,xB〉;〈x1,x4〉;〈xF,x6〉;〈x0,x3〉;〈xA,xB〉;〈xA,xE〉;〈xB,xC〉;〈xE,xB〉;〈xC,x8〉;〈x6,x7〉;〈xC,xC〉;〈xF,xF〉;〈x4,xF〉;〈xC,x6〉;〈x2,x9〉;〈x9,x5〉;
+〈xB,xC〉;〈x6,x5〉;〈x5,x2〉;〈xF,x2〉;〈x3,x5〉;〈xC,x4〉;〈xF,x4〉;〈x9,xB〉;〈x4,x5〉;〈x1,xC〉;〈xD,xB〉;〈x6,x1〉;〈xF,xE〉;〈x3,xF〉;〈xB,x9〉;〈xD,x8〉;
+〈xF,x1〉;〈xA,xC〉;〈x0,x7〉;〈xA,x4〉;〈xB,x8〉;〈x8,x8〉;〈x9,x5〉;〈xB,x8〉;〈x5,x6〉;〈x3,x2〉;〈x5,xA〉;〈x3,xE〉;〈x2,x2〉;〈x0,xB〉;〈x9,x6〉;〈xE,xE〉;
+〈x6,xF〉;〈x1,xE〉;〈x3,x2〉;〈x4,x9〉;〈x4,xF〉;〈xC,xC〉;〈xD,xB〉;〈x5,x1〉;〈x4,xD〉;〈xD,x1〉;〈x4,xF〉;〈x0,x9〉;〈x5,xA〉;〈xA,xC〉;〈xE,x7〉;〈x8,x6〉;
+〈x5,x8〉;〈xA,x5〉;〈xA,x7〉;〈x5,xE〉;〈x3,x6〉;〈x1,x1〉;〈x3,xA〉;〈x8,x9〉;〈x8,x2〉;〈xD,xC〉;〈x6,x2〉;〈x0,xE〉;〈xA,xB〉;〈x2,xB〉;〈x2,x5〉;〈xF,x9〉;
+〈x7,x7〉;〈x8,x6〉;〈x1,xD〉;〈x7,x9〉;〈x5,x1〉;〈xB,xD〉;〈x9,x8〉;〈xB,x7〉;〈xB,xB〉;〈xF,x6〉;〈xD,x9〉;〈x6,x6〉;〈x0,x1〉;〈x1,x2〉;〈xE,xB〉;〈x0,xA〉;
+〈xC,xD〉;〈x1,xA〉;〈xA,xA〉;〈xC,xC〉;〈x6,x5〉;〈x4,x2〉;〈x8,xF〉;〈x2,xA〉;〈x4,x8〉;〈xC,x6〉;〈xB,xA〉;〈xD,x8〉;〈x2,xD〉;〈x2,x9〉;〈xE,x8〉;〈x5,x7〉;
+〈x7,x7〉;〈x7,xA〉;〈xB,x4〉;〈x4,x9〉;〈x6,x5〉;〈x4,x3〉;〈x5,x7〉;〈xF,xE〉;〈xC,x6〉;〈xC,x7〉;〈x6,x2〉;〈x6,x7〉;〈x5,x8〉;〈xD,x6〉;〈x9,xA〉;〈xC,x8〉;
+〈xE,x8〉;〈x3,x0〉;〈x6,x0〉;〈x7,x3〉;〈x8,x9〉;〈x2,x3〉;〈x0,x8〉;〈x7,xA〉;〈xA,xC〉;〈x5,xD〉;〈x6,xD〉;〈xC,xE〉;〈x0,xC〉;〈x1,xB〉;〈x1,x7〉;〈xC,x1〉;
+〈x4,x2〉;〈x5,x3〉;〈x1,x5〉;〈x7,xC〉;〈x7,x4〉;〈x2,xB〉;〈x2,x5〉;〈x5,x6〉;〈x6,x1〉;〈xE,xC〉;〈x0,xB〉;〈x4,x2〉;〈x0,x4〉;〈xC,xA〉;〈x0,x9〉;〈xA,xB〉;
+〈x1,xB〉;〈xD,x0〉;〈x9,xF〉;〈x6,xA〉;〈x7,xF〉;〈x4,x1〉;〈xF,x8〉;〈xE,xA〉;〈x8,x2〉;〈x8,x1〉;〈x4,x1〉;〈xC,xE〉;〈xC,xE〉;〈x0,xD〉;〈x2,xB〉;〈x3,x3〉;
+〈xA,x3〉;〈x6,x4〉;〈xF,xA〉;〈xA,x6〉;〈x3,x9〉;〈x7,xF〉;〈xF,x6〉;〈xB,x2〉;〈x5,x5〉;〈x6,xB〉;〈xA,xC〉;〈x3,x3〉;〈x9,x3〉;〈xE,x7〉;〈xB,xE〉;〈x3,x4〉;
+〈xC,xF〉;〈xE,xF〉;〈xA,x2〉;〈xE,xE〉;〈xE,xD〉;〈xC,xB〉;〈xB,x0〉;〈x8,x9〉;〈xD,xA〉;〈x3,xB〉;〈xB,xE〉;〈x3,xE〉;〈x3,x3〉;〈x5,x1〉;〈xA,x5〉;〈x3,xC〉;
+〈xC,xC〉;〈xA,x0〉;〈xF,xD〉;〈x3,x9〉;〈xC,xB〉;〈xF,xC〉;〈x1,xF〉;〈x8,xD〉;〈x6,x8〉;〈xD,x4〉;〈x8,xC〉;〈xA,xA〉;〈x8,xE〉;〈x3,xA〉;〈x9,x7〉;〈x2,x6〉;
+〈x6,xB〉;〈xA,xC〉;〈x8,xA〉;〈x4,xB〉;〈x7,x4〉;〈x3,xF〉;〈xB,x7〉;〈xB,xF〉;〈x0,xC〉;〈xE,x6〉;〈xC,xD〉;〈x4,x2〉;〈xF,xA〉;〈xE,xE〉;〈xF,x9〉;〈x0,xC〉;
+〈x2,xC〉;〈x7,x9〉;〈x7,xE〉;〈xD,x8〉;〈x4,x0〉;〈x7,xC〉;〈x3,x8〉;〈x4,x9〉;〈x7,x1〉;〈x7,x5〉;〈xB,x7〉;〈x3,x6〉;〈x0,x7〉;〈x1,xA〉;〈x2,xC〉;〈x1,xE〉
+].
+
+definition dTest_random_ex10 ≝
+[
+〈x3,xC〉;〈x7,xA〉;〈x3,x8〉;〈x4,xA〉;〈x3,x4〉;〈x2,x0〉;〈x9,x5〉;〈x6,x0〉;〈xF,x7〉;〈xC,x3〉;〈xB,x1〉;〈x6,xE〉;〈xB,x1〉;〈x7,x0〉;〈x7,x4〉;〈x3,xB〉;
+〈x0,xD〉;〈x6,xD〉;〈xF,xB〉;〈xE,x5〉;〈xE,x2〉;〈x6,x6〉;〈x6,x8〉;〈x0,x8〉;〈xF,xB〉;〈x3,xC〉;〈x8,xC〉;〈xD,xD〉;〈x0,x2〉;〈x2,xE〉;〈x6,xE〉;〈xF,x1〉;
+〈xA,xF〉;〈x7,x9〉;〈x6,x7〉;〈xE,x2〉;〈x4,xC〉;〈xA,x5〉;〈x7,x9〉;〈xC,x6〉;〈xB,x5〉;〈xA,xF〉;〈x1,x5〉;〈xF,xE〉;〈xE,x2〉;〈x2,xB〉;〈xC,xA〉;〈xE,x6〉;
+〈x3,xE〉;〈x2,xC〉;〈x5,x8〉;〈x7,x2〉;〈xC,xE〉;〈x7,x1〉;〈x8,xC〉;〈xB,xE〉;〈x2,x0〉;〈x6,x6〉;〈x0,x7〉;〈x6,xF〉;〈xD,x1〉;〈x8,x2〉;〈x3,x1〉;〈xF,x3〉;
+〈x9,x5〉;〈x9,x1〉;〈x1,x2〉;〈xF,x3〉;〈x4,xF〉;〈x6,xC〉;〈xA,x6〉;〈x8,xE〉;〈xB,x2〉;〈x7,x8〉;〈xD,xE〉;〈x7,x9〉;〈xC,x5〉;〈x2,x2〉;〈xF,x3〉;〈x0,x7〉;
+〈xE,x7〉;〈x9,xE〉;〈x9,x2〉;〈x7,x3〉;〈x3,xC〉;〈xA,x1〉;〈xD,xA〉;〈x2,x1〉;〈x2,x3〉;〈x4,x5〉;〈xE,x5〉;〈x7,x4〉;〈x8,x4〉;〈xC,x2〉;〈x6,x8〉;〈x3,x5〉;
+〈x9,xA〉;〈xC,x8〉;〈x2,xE〉;〈x1,xD〉;〈xD,x1〉;〈xA,x6〉;〈xD,xF〉;〈x0,x6〉;〈x7,xF〉;〈x8,xC〉;〈x2,x8〉;〈xF,x6〉;〈xC,x3〉;〈xF,x8〉;〈x6,x2〉;〈xF,x9〉;
+〈x5,x7〉;〈x2,x7〉;〈xD,x1〉;〈xD,x3〉;〈x0,xB〉;〈xA,x3〉;〈x8,x7〉;〈x8,x3〉;〈xC,x9〉;〈x1,x4〉;〈xB,x4〉;〈xC,x5〉;〈xE,xD〉;〈x5,x4〉;〈xE,x1〉;〈xB,x9〉;
+〈x2,x9〉;〈x6,xF〉;〈xE,x3〉;〈x0,xD〉;〈xC,xC〉;〈xF,x6〉;〈x0,xD〉;〈x2,x4〉;〈x4,x5〉;〈x6,xE〉;〈xE,x4〉;〈xD,xB〉;〈xF,x9〉;〈xC,x1〉;〈xD,x4〉;〈xB,x4〉;
+〈xB,x5〉;〈x6,x6〉;〈x1,xC〉;〈x6,xE〉;〈xA,xF〉;〈x4,x0〉;〈xE,x6〉;〈x4,x9〉;〈x7,xE〉;〈x4,x1〉;〈x1,x8〉;〈x8,x7〉;〈xD,xF〉;〈xB,xB〉;〈x6,x0〉;〈x0,x5〉;
+〈xF,x4〉;〈x5,xD〉;〈xA,x1〉;〈xD,x6〉;〈x6,x7〉;〈x2,xA〉;〈xC,x8〉;〈x7,x7〉;〈xF,x9〉;〈x8,xA〉;〈xF,x9〉;〈x2,x6〉;〈xE,xF〉;〈x7,x4〉;〈x5,x8〉;〈x6,xA〉;
+〈xC,x8〉;〈x3,x5〉;〈x1,x0〉;〈xC,x5〉;〈x1,xE〉;〈x0,xB〉;〈x8,x3〉;〈x6,xA〉;〈x4,x4〉;〈x8,xD〉;〈x5,xC〉;〈xF,xB〉;〈xF,xE〉;〈x9,x2〉;〈x0,x3〉;〈x4,x3〉;
+〈x0,x5〉;〈xA,xF〉;〈xC,x0〉;〈xF,x3〉;〈x0,x4〉;〈x2,x8〉;〈x9,xD〉;〈x0,x9〉;〈x3,x5〉;〈xE,x3〉;〈x5,xF〉;〈x4,x5〉;〈xA,xB〉;〈xC,xD〉;〈x8,xC〉;〈xF,xD〉;
+〈x2,xC〉;〈x9,xD〉;〈xA,xF〉;〈x6,x4〉;〈x4,x3〉;〈x8,x0〉;〈x8,x2〉;〈xE,x5〉;〈x8,xE〉;〈x3,xD〉;〈x2,xD〉;〈xD,xB〉;〈xD,xF〉;〈xA,xB〉;〈x0,x8〉;〈x1,x6〉;
+〈xE,xC〉;〈x7,xE〉;〈xA,x7〉;〈xC,xB〉;〈xD,x8〉;〈x5,xC〉;〈x2,xC〉;〈x8,x8〉;〈x9,x8〉;〈xC,x2〉;〈xA,xD〉;〈x1,xD〉;〈xB,x0〉;〈xB,x1〉;〈xC,xE〉;〈x9,x3〉;
+〈xE,x2〉;〈xF,x4〉;〈xD,xB〉;〈xA,x5〉;〈xB,x6〉;〈x4,x9〉;〈x8,x7〉;〈x1,xD〉;〈xA,x2〉;〈x7,x9〉;〈x3,x5〉;〈xB,xE〉;〈x5,x5〉;〈xC,xD〉;〈x6,x3〉;〈x2,xC〉
+].
+
+definition dTest_random_ex11 ≝
+[
+〈x1,x0〉;〈x1,xF〉;〈xE,xC〉;〈x3,xB〉;〈x8,xA〉;〈x3,xF〉;〈x3,x8〉;〈x8,x0〉;〈x1,xC〉;〈x2,xD〉;〈x9,x2〉;〈x5,xF〉;〈xE,x1〉;〈xB,x5〉;〈xB,xC〉;〈x8,x3〉;
+〈xB,x6〉;〈x1,xB〉;〈xE,xD〉;〈x4,xF〉;〈x3,xA〉;〈xC,x4〉;〈xF,xE〉;〈xF,xF〉;〈xC,xB〉;〈x8,x1〉;〈x6,x7〉;〈xC,x2〉;〈x5,x9〉;〈xD,xA〉;〈x0,xA〉;〈x9,xC〉;
+〈x2,x2〉;〈xE,xB〉;〈x9,x3〉;〈xE,x2〉;〈x7,xF〉;〈xA,xC〉;〈x4,xA〉;〈x8,x2〉;〈x8,x1〉;〈x3,xF〉;〈xE,xB〉;〈x8,xB〉;〈x0,xF〉;〈x9,xC〉;〈x4,x1〉;〈x8,x2〉;
+〈x9,xB〉;〈x7,xC〉;〈x5,x1〉;〈xA,x7〉;〈xA,xB〉;〈xA,xD〉;〈x9,x2〉;〈x1,x9〉;〈xF,x0〉;〈xF,xD〉;〈x9,x3〉;〈xF,x6〉;〈xA,xD〉;〈x2,x4〉;〈xC,xB〉;〈xD,xE〉;
+〈xB,x5〉;〈xA,xB〉;〈x8,x1〉;〈x5,x4〉;〈xA,xE〉;〈x2,x4〉;〈x6,x4〉;〈xD,x2〉;〈xD,x0〉;〈xF,xE〉;〈x3,x3〉;〈x2,xA〉;〈x7,x5〉;〈x0,x7〉;〈x8,xF〉;〈x3,xA〉;
+〈x1,x2〉;〈x9,xF〉;〈xB,xE〉;〈x1,xB〉;〈x1,xB〉;〈x1,xF〉;〈xC,x7〉;〈xF,x1〉;〈x7,xC〉;〈x9,x1〉;〈x5,xD〉;〈x3,x2〉;〈xD,x9〉;〈xD,x6〉;〈xE,xA〉;〈x0,x6〉;
+〈x5,xB〉;〈x6,x9〉;〈x6,xB〉;〈xA,xC〉;〈x0,x9〉;〈x1,x6〉;〈xA,x3〉;〈xC,xA〉;〈x8,xE〉;〈x8,x3〉;〈x3,x4〉;〈xB,x7〉;〈x4,x1〉;〈x2,x7〉;〈xB,x3〉;〈x0,x1〉;
+〈xE,x6〉;〈x0,x6〉;〈x8,xC〉;〈x0,x4〉;〈x3,xD〉;〈xB,xE〉;〈x2,xC〉;〈x6,x6〉;〈xB,x5〉;〈x8,x6〉;〈x1,x1〉;〈x1,x3〉;〈x6,xD〉;〈xD,x0〉;〈xB,xE〉;〈x8,xD〉;
+〈xC,x7〉;〈x5,x5〉;〈x0,x2〉;〈xC,x1〉;〈x7,x6〉;〈x5,xF〉;〈x2,x0〉;〈x5,xE〉;〈xE,x4〉;〈x3,xE〉;〈x7,x7〉;〈xE,x1〉;〈x3,xF〉;〈xE,x8〉;〈x6,xC〉;〈x4,xA〉;
+〈xA,x0〉;〈xF,xE〉;〈xC,xE〉;〈x3,xF〉;〈x6,x7〉;〈x9,x4〉;〈x3,xF〉;〈xE,xF〉;〈xE,xF〉;〈x8,x6〉;〈xD,x9〉;〈x4,xA〉;〈x0,x8〉;〈x8,xB〉;〈xC,x8〉;〈x1,xC〉;
+〈xA,xD〉;〈x2,x0〉;〈xA,x7〉;〈x8,xC〉;〈x0,x6〉;〈x6,x7〉;〈xA,xF〉;〈x7,x3〉;〈xC,xD〉;〈x1,x6〉;〈x8,x4〉;〈x3,x2〉;〈xD,x0〉;〈xF,x3〉;〈xD,xC〉;〈xD,xB〉;
+〈xB,x7〉;〈x2,x4〉;〈x6,xA〉;〈x6,x3〉;〈x1,xC〉;〈xA,x1〉;〈xD,xE〉;〈xB,xC〉;〈x9,x2〉;〈xF,x1〉;〈x5,xC〉;〈xE,x7〉;〈xE,x0〉;〈xD,x5〉;〈xA,x4〉;〈x4,xA〉;
+〈x0,x0〉;〈xD,x6〉;〈x2,x2〉;〈x9,xC〉;〈x5,x2〉;〈x8,xF〉;〈xE,x8〉;〈x2,x2〉;〈xA,x2〉;〈xF,x0〉;〈x9,x8〉;〈x3,x8〉;〈x0,xD〉;〈xF,x6〉;〈x4,x3〉;〈x7,x9〉;
+〈x8,x2〉;〈xA,xF〉;〈xD,x5〉;〈xC,x1〉;〈x8,x2〉;〈x5,x2〉;〈xD,xB〉;〈x8,xF〉;〈x7,xE〉;〈xD,x1〉;〈x9,xD〉;〈xA,x6〉;〈x8,xE〉;〈x9,xE〉;〈xA,x9〉;〈x8,xE〉;
+〈xD,x1〉;〈xF,x6〉;〈xB,x0〉;〈xE,xA〉;〈x8,x3〉;〈xE,x9〉;〈xF,x7〉;〈x3,xB〉;〈x4,xA〉;〈x0,x9〉;〈x1,xE〉;〈x3,x2〉;〈xD,x2〉;〈x5,xD〉;〈xD,x7〉;〈xA,xB〉;
+〈x4,xD〉;〈x6,xF〉;〈x5,x9〉;〈xF,xC〉;〈x4,x3〉;〈x4,x1〉;〈x0,x0〉;〈x3,xC〉;〈x9,x4〉;〈x5,x2〉;〈x5,x9〉;〈x6,xC〉;〈x6,xE〉;〈xE,x8〉;〈x6,x6〉;〈xF,x5〉
+].
+
+definition dTest_random_ex12 ≝
+[
+〈x9,xC〉;〈x5,x7〉;〈x6,xC〉;〈xE,x2〉;〈x3,xB〉;〈xA,x2〉;〈x2,x1〉;〈xE,xE〉;〈xF,x6〉;〈x4,xF〉;〈xF,x3〉;〈x6,x2〉;〈xD,xB〉;〈x8,xF〉;〈x6,x4〉;〈xD,x3〉;
+〈x8,x0〉;〈x6,x9〉;〈x9,x7〉;〈x4,x7〉;〈x8,xB〉;〈xB,x6〉;〈x3,x8〉;〈x4,x5〉;〈xB,xE〉;〈x0,xD〉;〈x6,x1〉;〈xC,xF〉;〈x7,x8〉;〈xC,xF〉;〈x4,x1〉;〈x7,xF〉;
+〈xF,x4〉;〈x5,xA〉;〈x8,xB〉;〈x7,x2〉;〈xE,x6〉;〈x7,xD〉;〈x4,xC〉;〈x1,x8〉;〈xE,xE〉;〈x3,xA〉;〈x1,x2〉;〈x9,x4〉;〈x3,x4〉;〈x3,x0〉;〈x3,x9〉;〈x0,x0〉;
+〈x9,x5〉;〈x6,x0〉;〈xF,xA〉;〈x7,xF〉;〈xA,x6〉;〈xC,x7〉;〈xB,x1〉;〈x7,xE〉;〈x0,xD〉;〈x2,x4〉;〈xF,xF〉;〈x4,x3〉;〈x7,x8〉;〈x8,x8〉;〈x6,xC〉;〈x0,x7〉;
+〈x7,x3〉;〈x9,x2〉;〈xC,x8〉;〈x0,xB〉;〈x5,x0〉;〈x9,x7〉;〈xF,x4〉;〈x1,xB〉;〈xD,xB〉;〈x4,x5〉;〈x6,x2〉;〈x9,x1〉;〈x8,x7〉;〈x5,xD〉;〈xF,x5〉;〈x6,x1〉;
+〈x3,x8〉;〈xF,x3〉;〈x8,xA〉;〈x4,xF〉;〈xD,xB〉;〈x3,xD〉;〈x4,x3〉;〈x2,xF〉;〈xB,xA〉;〈xA,x9〉;〈xB,x4〉;〈xA,x9〉;〈x2,x6〉;〈x7,xE〉;〈x8,xC〉;〈x1,x6〉;
+〈xF,xC〉;〈x1,xC〉;〈xA,x7〉;〈x9,xD〉;〈x7,xC〉;〈x8,x5〉;〈xA,x7〉;〈x3,x5〉;〈x8,x6〉;〈x4,xF〉;〈xC,x8〉;〈x9,xA〉;〈x0,xD〉;〈x2,x4〉;〈x0,x4〉;〈x5,x8〉;
+〈x4,x2〉;〈x3,x0〉;〈x7,x5〉;〈x8,xD〉;〈xB,xC〉;〈x7,x2〉;〈x3,x4〉;〈xF,x9〉;〈x4,x6〉;〈x5,xE〉;〈xF,x4〉;〈x5,xC〉;〈x6,x7〉;〈x0,x4〉;〈x8,x8〉;〈x2,x8〉;
+〈xA,x9〉;〈xE,xC〉;〈xE,xC〉;〈x7,xC〉;〈x4,x6〉;〈x6,xE〉;〈x1,xC〉;〈xD,x9〉;〈xC,x0〉;〈x6,x1〉;〈x5,x8〉;〈x5,xE〉;〈x3,x7〉;〈xB,x0〉;〈x9,x9〉;〈x9,x7〉;
+〈xA,x0〉;〈x9,xB〉;〈x7,x7〉;〈x4,xA〉;〈xD,x0〉;〈xB,x5〉;〈x1,xD〉;〈xA,x6〉;〈x4,xA〉;〈xC,x5〉;〈x3,xA〉;〈x9,x4〉;〈xB,x4〉;〈x5,x0〉;〈x0,x6〉;〈xC,x0〉;
+〈xA,x8〉;〈x0,x2〉;〈x5,xF〉;〈x0,xE〉;〈x2,x1〉;〈x0,x3〉;〈xB,x1〉;〈x9,x6〉;〈x0,x2〉;〈x9,x7〉;〈x8,x0〉;〈x1,xA〉;〈x0,xB〉;〈x3,xD〉;〈x2,x1〉;〈x7,xF〉;
+〈x0,x3〉;〈x2,x9〉;〈x3,x2〉;〈x6,x6〉;〈xB,xF〉;〈x3,xB〉;〈x5,x7〉;〈x5,x3〉;〈xE,x7〉;〈xD,x5〉;〈xE,x5〉;〈x4,x5〉;〈xA,x3〉;〈x1,xA〉;〈x1,xB〉;〈xF,x8〉;
+〈xC,xF〉;〈xB,x3〉;〈x9,x5〉;〈x5,x9〉;〈x4,xE〉;〈x6,x4〉;〈x4,x3〉;〈xF,x4〉;〈x2,x5〉;〈xC,xC〉;〈x6,x1〉;〈x3,x5〉;〈xD,xF〉;〈x3,x6〉;〈x5,x5〉;〈xC,xF〉;
+〈x9,xA〉;〈x1,x1〉;〈xF,x6〉;〈xD,x4〉;〈x4,xF〉;〈x9,xB〉;〈xA,xF〉;〈xF,x2〉;〈x0,x3〉;〈x1,x9〉;〈x9,xB〉;〈xA,xB〉;〈xC,x4〉;〈x1,x9〉;〈xA,x1〉;〈xE,xA〉;
+〈x1,xB〉;〈x2,x5〉;〈xA,xD〉;〈xA,xA〉;〈x0,x0〉;〈x5,xB〉;〈x9,xD〉;〈x6,xF〉;〈x8,x8〉;〈x6,xF〉;〈x3,x0〉;〈x8,x5〉;〈xC,x6〉;〈x1,x7〉;〈x5,x7〉;〈x1,x1〉;
+〈xA,xB〉;〈x0,x2〉;〈xD,xD〉;〈x9,x2〉;〈x4,xD〉;〈x8,x2〉;〈x0,x2〉;〈x3,x5〉;〈xC,xB〉;〈x4,x4〉;〈xA,x4〉;〈x4,x1〉;〈xD,x5〉;〈x1,x2〉;〈xE,x7〉;〈x4,xD〉
+].
+
+definition dTest_random_32 ≝ 
+ dTest_random_ex01a.
+
+definition dTest_random_64 ≝ 
+ dTest_random_32@dTest_random_ex01b.
+
+definition dTest_random_128 ≝ 
+ dTest_random_64@dTest_random_ex01c.
+
+definition dTest_random_256 ≝ 
+ dTest_random_128@dTest_random_ex01d.
+
+definition dTest_random_512 ≝ 
+ dTest_random_256@dTest_random_ex02.
+
+definition dTest_random_1024 ≝ 
+ dTest_random_512@dTest_random_ex03@dTest_random_ex04.
+
+definition dTest_random_2048 ≝ 
+ dTest_random_1024@dTest_random_ex05@dTest_random_ex06@dTest_random_ex07@dTest_random_ex08.
+
+definition dTest_random_3072 ≝ 
+ dTest_random_2048@dTest_random_ex09@dTest_random_ex10@dTest_random_ex11@dTest_random_ex12.
+
+(* campitura di 0x00 *)
+definition dTest_bytes_aux : list byte8 ≝
+[
+〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;
+〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;
+〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;
+〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;
+〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;
+〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;
+〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;
+〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉;〈x0,x0〉
+].
+
+(* blocco di 0x00 lungo 0x0380 da caricare dopo dati per azzerare
+   counters, index, position, e stack di esecuzione *)
+definition dTest_bytes : list byte8 ≝
+ dTest_bytes_aux@dTest_bytes_aux@dTest_bytes_aux@dTest_bytes_aux@
+ dTest_bytes_aux@dTest_bytes_aux@dTest_bytes_aux.
diff --git a/helm/software/matita/library/freescale/memory_abs.ma b/helm/software/matita/library/freescale/memory_abs.ma
new file mode 100644 (file)
index 0000000..23f3f15
--- /dev/null
@@ -0,0 +1,259 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/memory_abs/".
+
+(*include "/media/VIRTUOSO/freescale/memory_func.ma".*)
+include "freescale/memory_func.ma".
+(*include "/media/VIRTUOSO/freescale/memory_trees.ma".*)
+include "freescale/memory_trees.ma".
+(*include "/media/VIRTUOSO/freescale/memory_bits.ma".*)
+include "freescale/memory_bits.ma".
+
+(* ********************************************* *)
+(* ASTRAZIONE DALL'IMPLEMENTAZIONE DELLA MEMORIA *)
+(* ********************************************* *)
+
+(* tipi di implementazione della memoria *)
+inductive memory_impl : Type ≝
+  MEM_FUNC: memory_impl
+| MEM_TREE: memory_impl
+| MEM_BITS: memory_impl.
+
+(* ausiliario per il tipo della memoria *)
+definition aux_mem_type ≝
+λt:memory_impl.match t with
+ [ MEM_FUNC ⇒ word16 → byte8
+ | MEM_TREE ⇒ Prod16T (Prod16T (Prod16T (Prod16T byte8)))
+ | MEM_BITS ⇒ Prod16T (Prod16T (Prod16T (Prod16T (Prod8T bool))))
+ ].
+
+(* ausiliario per il tipo del checker *)
+definition aux_chk_type ≝
+λt:memory_impl.match t with
+ [ MEM_FUNC ⇒ word16 → memory_type
+ | MEM_TREE ⇒ Prod16T (Prod16T (Prod16T (Prod16T memory_type)))
+ | MEM_BITS ⇒ Prod16T (Prod16T (Prod16T (Prod16T (Prod8T memory_type))))
+ ].
+
+(* unificazione di out_of_bound_memory *)
+definition out_of_bound_memory ≝
+λt:memory_impl.
+ match t
+  return λt.aux_chk_type t
+ with
+  [ MEM_FUNC ⇒ mf_out_of_bound_memory
+  | MEM_TREE ⇒ mt_out_of_bound_memory
+  | MEM_BITS ⇒ mb_out_of_bound_memory
+  ].
+
+(* unificazione di zero_memory *)
+definition zero_memory ≝
+λt:memory_impl.
+ match t
+  return λt.aux_mem_type t
+ with
+  [ MEM_FUNC ⇒ mf_zero_memory
+  | MEM_TREE ⇒ mt_zero_memory
+  | MEM_BITS ⇒ mb_zero_memory
+  ].
+
+(* unificazione della lettura senza chk: mem_read_abs mem addr *)
+definition mem_read_abs ≝
+λt:memory_impl.
+ match t
+  return λt.aux_mem_type t → word16 → byte8 
+ with
+  [ MEM_FUNC ⇒ λm:aux_mem_type MEM_FUNC.
+               λaddr:word16.
+               m addr
+  | MEM_TREE ⇒ λm:aux_mem_type MEM_TREE.
+               λaddr:word16.
+               mt_visit byte8 m addr
+  | MEM_BITS ⇒ λm:aux_mem_type MEM_BITS.
+               λaddr:word16.
+               byte8_of_bits (mt_visit (Prod8T bool) m addr)
+  ].
+
+(* unificazione della lettura con chk: mem_read mem chk addr *)
+definition mem_read ≝
+λt:memory_impl.
+ match t
+  return λt.aux_mem_type t → aux_chk_type t → word16 → option byte8 
+ with
+  [ MEM_FUNC ⇒ λm:aux_mem_type MEM_FUNC.
+               λc:aux_chk_type MEM_FUNC.
+               λaddr:word16.
+               mf_mem_read m c addr
+  | MEM_TREE ⇒ λm:aux_mem_type MEM_TREE.
+               λc:aux_chk_type MEM_TREE.
+               λaddr:word16.
+               mt_mem_read m c addr
+  | MEM_BITS ⇒ λm:aux_mem_type MEM_BITS.
+               λc:aux_chk_type MEM_BITS.
+               λaddr:word16.
+               mb_mem_read m c addr
+  ].
+
+(* unificazione della lettura di bit con chk: mem_read mem chk addr sub *)
+definition mem_read_bit ≝
+λt:memory_impl.
+ match t
+  return λt.aux_mem_type t → aux_chk_type t → word16 → oct → option bool 
+ with
+  [ MEM_FUNC ⇒ λm:aux_mem_type MEM_FUNC.
+               λc:aux_chk_type MEM_FUNC.
+               λaddr:word16.
+               λo:oct.
+               opt_map ?? (mf_mem_read m c addr)
+                (λb.Some ? (getn_array8T o bool (bits_of_byte8 b))) 
+  | MEM_TREE ⇒ λm:aux_mem_type MEM_TREE.
+               λc:aux_chk_type MEM_TREE.
+               λaddr:word16.
+               λo:oct.
+               opt_map ?? (mt_mem_read m c addr)
+                (λb.Some ? (getn_array8T o bool (bits_of_byte8 b)))
+  | MEM_BITS ⇒ λm:aux_mem_type MEM_BITS.
+               λc:aux_chk_type MEM_BITS.
+               λaddr:word16.
+               λo:oct.
+               mb_mem_read_bit m c addr o
+  ].
+
+(* unificazione della scrittura con chk: mem_update mem chk addr val *)
+definition mem_update ≝
+λt:memory_impl.
+ match t
+  return λt.aux_mem_type t → aux_chk_type t → word16 → byte8 → option (aux_mem_type t) 
+ with
+  [ MEM_FUNC ⇒ λm:aux_mem_type MEM_FUNC.
+               λc:aux_chk_type MEM_FUNC.
+               λaddr:word16.
+               λv:byte8.
+               mf_mem_update m c addr v
+  | MEM_TREE ⇒ λm:aux_mem_type MEM_TREE.
+               λc:aux_chk_type MEM_TREE.
+               λaddr:word16.
+               λv:byte8.
+               mt_mem_update m c addr v
+  | MEM_BITS ⇒ λm:aux_mem_type MEM_BITS.
+               λc:aux_chk_type MEM_BITS.
+               λaddr:word16.
+               λv:byte8.
+               mb_mem_update m c addr v
+  ].
+
+(* unificazione della scrittura di bit con chk: mem_update mem chk addr sub val *)
+definition mem_update_bit ≝
+λt:memory_impl.
+ match t
+  return λt.aux_mem_type t → aux_chk_type t → word16 → oct → bool → option (aux_mem_type t) 
+ with
+  [ MEM_FUNC ⇒ λm:aux_mem_type MEM_FUNC.
+               λc:aux_chk_type MEM_FUNC.
+               λaddr:word16.
+               λo:oct.
+               λv:bool.
+               opt_map ?? (mf_mem_read m c addr)
+                (λb.mf_mem_update m c addr (byte8_of_bits (setn_array8T o bool (bits_of_byte8 b) v)))
+  | MEM_TREE ⇒ λm:aux_mem_type MEM_TREE.
+               λc:aux_chk_type MEM_TREE.
+               λaddr:word16.
+               λo:oct.
+               λv:bool.
+               opt_map ?? (mt_mem_read m c addr)
+                (λb.mt_mem_update m c addr (byte8_of_bits (setn_array8T o bool (bits_of_byte8 b) v)))
+  | MEM_BITS ⇒ λm:aux_mem_type MEM_BITS.
+               λc:aux_chk_type MEM_BITS.
+               λaddr:word16.
+               λo:oct.
+               λv:bool.
+               mb_mem_update_bit m c addr o v
+  ].
+
+(* unificazione del caricamento: load_from_source_at old_mem source addr *)
+definition load_from_source_at ≝
+λt:memory_impl.
+ match t
+  return λt.aux_mem_type t → list byte8 → word16 → aux_mem_type t 
+ with
+  [ MEM_FUNC ⇒ λm:aux_mem_type MEM_FUNC.
+               λl:list byte8.
+               λaddr:word16.
+               mf_load_from_source_at m l addr
+  | MEM_TREE ⇒ λm:aux_mem_type MEM_TREE.
+               λl:list byte8.
+               λaddr:word16.
+               mt_load_from_source_at m l addr
+  | MEM_BITS ⇒ λm:aux_mem_type MEM_BITS.
+               λl:list byte8.
+               λaddr:word16.
+               mb_load_from_source_at m l addr
+  ].
+
+(* unificazione dell'impostazione della memoria: chk_update_ranged chk inf sup v *)
+definition check_update_ranged ≝
+λt:memory_impl.
+ match t
+  return λt.aux_chk_type t → word16 → word16 → memory_type → aux_chk_type t 
+ with
+  [ MEM_FUNC ⇒ λc:aux_chk_type MEM_FUNC.
+               λinf,sup:word16.
+               λv:memory_type.
+               mf_check_update_ranged c inf sup v
+  | MEM_TREE ⇒ λc:aux_chk_type MEM_TREE.
+               λinf,sup:word16.
+               λv:memory_type.
+               mt_update_ranged memory_type c inf sup v
+  | MEM_BITS ⇒ λc:aux_chk_type MEM_BITS.
+               λinf,sup:word16.
+               λv:memory_type.
+               mt_update_ranged (Prod8T memory_type) c inf sup (array_8T memory_type v v v v v v v v)
+  ].
+
+(* unificazione dell'impostazione dei bit: chk_update_bit chk addr sub v *)
+(* NB: dove non esiste la granularita' del bit, lascio inalterato *)
+definition check_update_bit ≝
+λt:memory_impl.
+ match t
+  return λt.aux_chk_type t → word16 → oct → memory_type → aux_chk_type t
+ with
+  [ MEM_FUNC ⇒ λc:aux_chk_type MEM_FUNC.
+               λaddr:word16.
+               λo:oct.
+               λv:memory_type.
+               c
+  | MEM_TREE ⇒ λc:aux_chk_type MEM_TREE.
+               λaddr:word16.
+               λo:oct.
+               λv:memory_type.
+               c
+  | MEM_BITS ⇒ λc:aux_chk_type MEM_BITS.
+               λaddr:word16.
+               λo:oct.
+               λv:memory_type.
+               mb_chk_update_bit c addr o v
+  ].
diff --git a/helm/software/matita/library/freescale/memory_bits.ma b/helm/software/matita/library/freescale/memory_bits.ma
new file mode 100644 (file)
index 0000000..d7425bf
--- /dev/null
@@ -0,0 +1,227 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/memory_bits/".
+
+(*include "/media/VIRTUOSO/freescale/memory_trees.ma".*)
+include "freescale/memory_trees.ma".
+
+(* ********************* *)
+(* MEMORIA E DESCRITTORE *)
+(* ********************* *)
+
+(* tutta la memoria non installata *)
+definition mb_out_of_bound_memory ≝
+let base ≝ array_8T memory_type MEM_OUT_OF_BOUND MEM_OUT_OF_BOUND MEM_OUT_OF_BOUND MEM_OUT_OF_BOUND
+                                MEM_OUT_OF_BOUND MEM_OUT_OF_BOUND MEM_OUT_OF_BOUND MEM_OUT_OF_BOUND in
+let lev4 ≝ array_16T ? 
+           base base base base base base base base
+           base base base base base base base base
+           in
+let lev3 ≝ array_16T ?
+           lev4 lev4 lev4 lev4 lev4 lev4 lev4 lev4
+           lev4 lev4 lev4 lev4 lev4 lev4 lev4 lev4
+           in
+let lev2 ≝ array_16T ?
+           lev3 lev3 lev3 lev3 lev3 lev3 lev3 lev3
+           lev3 lev3 lev3 lev3 lev3 lev3 lev3 lev3
+           in
+let lev1 ≝ array_16T ?
+           lev2 lev2 lev2 lev2 lev2 lev2 lev2 lev2
+           lev2 lev2 lev2 lev2 lev2 lev2 lev2 lev2
+           in
+lev1.
+
+(* tutta la memoria a 0 *)
+definition mb_zero_memory ≝
+let base ≝ array_8T bool false false false false false false false false in
+let lev4 ≝ array_16T ? 
+           base base base base base base base base
+           base base base base base base base base
+           in
+let lev3 ≝ array_16T ?
+           lev4 lev4 lev4 lev4 lev4 lev4 lev4 lev4
+           lev4 lev4 lev4 lev4 lev4 lev4 lev4 lev4
+           in
+let lev2 ≝ array_16T ?
+           lev3 lev3 lev3 lev3 lev3 lev3 lev3 lev3
+           lev3 lev3 lev3 lev3 lev3 lev3 lev3 lev3
+           in
+let lev1 ≝ array_16T ?
+           lev2 lev2 lev2 lev2 lev2 lev2 lev2 lev2
+           lev2 lev2 lev2 lev2 lev2 lev2 lev2 lev2
+           in
+lev1.
+
+(* scrivi bit controllando il tipo di memoria *)
+definition mb_mem_update_bit ≝
+λmem:Prod16T (Prod16T (Prod16T (Prod16T (Prod8T bool)))).
+λchk:Prod16T (Prod16T (Prod16T (Prod16T (Prod8T memory_type)))).
+λaddr:word16.λsub:oct.λv:bool.
+ match getn_array8T sub memory_type (mt_visit (Prod8T memory_type) chk addr) with
+  (* ROM? ok, ma il valore viene perso *)
+  [ MEM_READ_ONLY ⇒ Some ? mem
+  (* RAM? ok *)
+  | MEM_READ_WRITE ⇒ Some ? (mt_update (Prod8T bool) mem addr (setn_array8T sub bool (mt_visit (Prod8T bool) mem addr) v))
+  (* NON INSTALLATA? no *)
+  | MEM_OUT_OF_BOUND ⇒ None ? ].
+
+(* scrivi tipo di bit *)
+definition mb_chk_update_bit ≝
+λchk:Prod16T (Prod16T (Prod16T (Prod16T (Prod8T memory_type)))).
+λaddr:word16.λsub:oct.λv:memory_type.
+ mt_update (Prod8T memory_type) chk addr (setn_array8T sub memory_type (mt_visit (Prod8T memory_type) chk addr) v).
+
+(* leggi bit controllando il tipo di memoria *)
+definition mb_mem_read_bit ≝
+λmem:Prod16T (Prod16T (Prod16T (Prod16T (Prod8T bool)))).
+λchk:Prod16T (Prod16T (Prod16T (Prod16T (Prod8T memory_type)))).
+λaddr:word16.λsub:oct.
+ match getn_array8T sub memory_type (mt_visit (Prod8T memory_type) chk addr) with
+  (* ROM? ok, ma il valore viene perso *)
+  [ MEM_READ_ONLY ⇒ Some ? (getn_array8T sub bool (mt_visit (Prod8T bool) mem addr))
+  (* RAM? ok *)
+  | MEM_READ_WRITE ⇒ Some ? (getn_array8T sub bool (mt_visit (Prod8T bool) mem addr))
+  (* NON INSTALLATA? no *)
+  | MEM_OUT_OF_BOUND ⇒ None ? ]. 
+
+(* scrivi controllando il tipo di memoria *)
+(* NB: devono esistere tutti i bit *)
+definition mb_mem_update ≝
+λmem:Prod16T (Prod16T (Prod16T (Prod16T (Prod8T bool)))).
+λchk:Prod16T (Prod16T (Prod16T (Prod16T (Prod8T memory_type)))).
+λaddr:word16.λv:byte8.
+let bit_types ≝ mt_visit (Prod8T memory_type) chk addr in
+let old_value ≝ mt_visit (Prod8T bool) mem addr in
+let new_value ≝ bits_of_byte8 v in
+let newbit0 ≝ match getn_array8T o0 memory_type bit_types with
+ [ MEM_READ_ONLY ⇒ Some bool (getn_array8T o0 bool old_value)
+ | MEM_READ_WRITE ⇒ Some bool (getn_array8T o0 bool new_value)
+ | MEM_OUT_OF_BOUND ⇒ None bool ] in
+let newbit1 ≝ match getn_array8T o1 memory_type bit_types with
+ [ MEM_READ_ONLY ⇒ Some bool (getn_array8T o1 bool old_value)
+ | MEM_READ_WRITE ⇒ Some bool (getn_array8T o1 bool new_value)
+ | MEM_OUT_OF_BOUND ⇒ None bool ] in
+let newbit2 ≝ match getn_array8T o2 memory_type bit_types with
+ [ MEM_READ_ONLY ⇒ Some bool (getn_array8T o2 bool old_value)
+ | MEM_READ_WRITE ⇒ Some bool (getn_array8T o2 bool new_value)
+ | MEM_OUT_OF_BOUND ⇒ None bool ] in
+let newbit3 ≝ match getn_array8T o3 memory_type bit_types with
+ [ MEM_READ_ONLY ⇒ Some bool (getn_array8T o3 bool old_value)
+ | MEM_READ_WRITE ⇒ Some bool (getn_array8T o3 bool new_value)
+ | MEM_OUT_OF_BOUND ⇒ None bool ] in
+let newbit4 ≝ match getn_array8T o4 memory_type bit_types with
+ [ MEM_READ_ONLY ⇒ Some bool (getn_array8T o4 bool old_value)
+ | MEM_READ_WRITE ⇒ Some bool (getn_array8T o4 bool new_value)
+ | MEM_OUT_OF_BOUND ⇒ None bool ] in
+let newbit5 ≝ match getn_array8T o5 memory_type bit_types with
+ [ MEM_READ_ONLY ⇒ Some bool (getn_array8T o5 bool old_value)
+ | MEM_READ_WRITE ⇒ Some bool (getn_array8T o5 bool new_value)
+ | MEM_OUT_OF_BOUND ⇒ None bool ] in
+let newbit6 ≝ match getn_array8T o6 memory_type bit_types with
+ [ MEM_READ_ONLY ⇒ Some bool (getn_array8T o6 bool old_value)
+ | MEM_READ_WRITE ⇒ Some bool (getn_array8T o6 bool new_value)
+ | MEM_OUT_OF_BOUND ⇒ None bool ] in
+let newbit7 ≝ match getn_array8T o7 memory_type bit_types with
+ [ MEM_READ_ONLY ⇒ Some bool (getn_array8T o7 bool old_value)
+ | MEM_READ_WRITE ⇒ Some bool (getn_array8T o7 bool new_value)
+ | MEM_OUT_OF_BOUND ⇒ None bool ] in
+       opt_map ?? newbit0
+ (λnb0.opt_map ?? newbit1
+ (λnb1.opt_map ?? newbit2
+ (λnb2.opt_map ?? newbit3
+ (λnb3.opt_map ?? newbit4
+ (λnb4.opt_map ?? newbit5
+ (λnb5.opt_map ?? newbit6
+ (λnb6.opt_map ?? newbit7
+ (λnb7.Some ? (mt_update (Prod8T bool) mem addr (array_8T bool nb7 nb6 nb5 nb4 nb3 nb2 nb1 nb0)))))))))).
+
+(* leggi controllando il tipo di memoria *)
+(* NB: devono esistere tutti i bit *)
+definition mb_mem_read ≝
+λmem:Prod16T (Prod16T (Prod16T (Prod16T (Prod8T bool)))).
+λchk:Prod16T (Prod16T (Prod16T (Prod16T (Prod8T memory_type)))).
+λaddr:word16.
+let bit_types ≝ mt_visit (Prod8T memory_type) chk addr in
+let value ≝ mt_visit (Prod8T bool) mem addr in
+let newbit0 ≝ match getn_array8T o0 memory_type bit_types with
+ [ MEM_READ_ONLY ⇒ Some bool (getn_array8T o0 bool value)
+ | MEM_READ_WRITE ⇒ Some bool (getn_array8T o0 bool value)
+ | MEM_OUT_OF_BOUND ⇒ None bool ] in
+let newbit1 ≝ match getn_array8T o1 memory_type bit_types with
+ [ MEM_READ_ONLY ⇒ Some bool (getn_array8T o1 bool value)
+ | MEM_READ_WRITE ⇒ Some bool (getn_array8T o1 bool value)
+ | MEM_OUT_OF_BOUND ⇒ None bool ] in
+let newbit2 ≝ match getn_array8T o2 memory_type bit_types with
+ [ MEM_READ_ONLY ⇒ Some bool (getn_array8T o2 bool value)
+ | MEM_READ_WRITE ⇒ Some bool (getn_array8T o2 bool value)
+ | MEM_OUT_OF_BOUND ⇒ None bool ] in
+let newbit3 ≝ match getn_array8T o3 memory_type bit_types with
+ [ MEM_READ_ONLY ⇒ Some bool (getn_array8T o3 bool value)
+ | MEM_READ_WRITE ⇒ Some bool (getn_array8T o3 bool value)
+ | MEM_OUT_OF_BOUND ⇒ None bool ] in
+let newbit4 ≝ match getn_array8T o4 memory_type bit_types with
+ [ MEM_READ_ONLY ⇒ Some bool (getn_array8T o4 bool value)
+ | MEM_READ_WRITE ⇒ Some bool (getn_array8T o4 bool value)
+ | MEM_OUT_OF_BOUND ⇒ None bool ] in
+let newbit5 ≝ match getn_array8T o5 memory_type bit_types with
+ [ MEM_READ_ONLY ⇒ Some bool (getn_array8T o5 bool value)
+ | MEM_READ_WRITE ⇒ Some bool (getn_array8T o5 bool value)
+ | MEM_OUT_OF_BOUND ⇒ None bool ] in
+let newbit6 ≝ match getn_array8T o6 memory_type bit_types with
+ [ MEM_READ_ONLY ⇒ Some bool (getn_array8T o6 bool value)
+ | MEM_READ_WRITE ⇒ Some bool (getn_array8T o6 bool value)
+ | MEM_OUT_OF_BOUND ⇒ None bool ] in
+let newbit7 ≝ match getn_array8T o7 memory_type bit_types with
+ [ MEM_READ_ONLY ⇒ Some bool (getn_array8T o7 bool value)
+ | MEM_READ_WRITE ⇒ Some bool (getn_array8T o7 bool value)
+ | MEM_OUT_OF_BOUND ⇒ None bool ] in
+       opt_map ?? newbit0
+ (λnb0.opt_map ?? newbit1
+ (λnb1.opt_map ?? newbit2
+ (λnb2.opt_map ?? newbit3
+ (λnb3.opt_map ?? newbit4
+ (λnb4.opt_map ?? newbit5
+ (λnb5.opt_map ?? newbit6
+ (λnb6.opt_map ?? newbit7
+ (λnb7.Some ? (byte8_of_bits (array_8T bool nb7 nb6 nb5 nb4 nb3 nb2 nb1 nb0)))))))))).
+
+(* ************************** *)
+(* CARICAMENTO PROGRAMMA/DATI *)
+(* ************************** *)
+
+(* carica a paratire da addr, scartando source (pescando da old_mem) se si supera 0xFFFF... *)
+let rec mb_load_from_source_at (old_mem:Prod16T (Prod16T (Prod16T (Prod16T (Prod8T bool)))))
+                               (source:list byte8) (addr:word16) on source ≝
+match source with
+ (* fine di source: carica da old_mem *)
+ [ nil ⇒ old_mem
+ | cons hd tl ⇒ match lt_w16 addr 〈〈xF,xF〉:〈xF,xF〉〉 with
+  (* non supera 0xFFFF, ricorsione *)
+  [ true ⇒ mb_load_from_source_at (mt_update ? old_mem addr (bits_of_byte8 hd)) tl (plus_w16nc addr 〈〈x0,x0〉:〈x0,x1〉〉)
+  (* supera 0xFFFF, niente ricorsione *)
+  | false ⇒ mt_update ? old_mem addr (bits_of_byte8 hd)
+  ]].
diff --git a/helm/software/matita/library/freescale/memory_func.ma b/helm/software/matita/library/freescale/memory_func.ma
new file mode 100644 (file)
index 0000000..44b2524
--- /dev/null
@@ -0,0 +1,153 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/memory_func/".
+
+(*include "/media/VIRTUOSO/freescale/memory_struct.ma".*)
+include "freescale/memory_struct.ma".
+
+(* ********************* *)
+(* MEMORIA E DESCRITTORE *)
+(* ********************* *)
+
+(* (mf_check_update_ranged chk inf sup mode) = setta tipo memoria *)
+definition mf_check_update_ranged ≝
+λf:word16 → memory_type.λi.λs.λv.
+ λx.match in_range x i s with
+  [ true ⇒ v
+  | false ⇒ f x ].
+
+(* tutta la memoria non installata *)
+definition mf_out_of_bound_memory ≝ λ_:word16.MEM_OUT_OF_BOUND.
+
+(* (mf_mem_update mem check addr val) = scrivi controllando il tipo di memoria *)
+definition mf_mem_update ≝
+λf:word16 → byte8.λc:word16 → memory_type.λa.λv.
+ match c a with
+  (* ROM? ok, ma il valore viene perso *)
+  [ MEM_READ_ONLY ⇒ Some ? f
+  (* RAM? ok *)
+  | MEM_READ_WRITE ⇒ Some ? (λx.match eq_w16 x a with [ true ⇒ v | false ⇒ f x ])
+  (* NON INSTALLATA? no *)
+  | MEM_OUT_OF_BOUND ⇒ None ? ].  
+
+(* tutta la memoria a 0 *)
+definition mf_zero_memory ≝ λ_:word16.〈x0,x0〉.
+
+(* (mf_mem_read mem check addr) = leggi controllando il tipo di memoria *)
+definition mf_mem_read ≝
+λf:word16 → byte8.λc:word16 → memory_type.λa.
+ match c a with
+  [ MEM_READ_ONLY ⇒ Some ? (f a)
+  | MEM_READ_WRITE ⇒ Some ? (f a)
+  | MEM_OUT_OF_BOUND ⇒ None ? ].
+
+(* ************************** *)
+(* CARICAMENTO PROGRAMMA/DATI *)
+(* ************************** *)
+
+(* carica a paratire da addr, scartando source (pescando da old_mem) se si supera 0xFFFF... *)
+let rec mf_load_from_source_at (old_mem:word16 → byte8) (source:list byte8) (addr:word16) on source ≝
+match source with
+ (* fine di source: carica da old_mem *)
+ [ nil ⇒ old_mem
+ | cons hd tl ⇒ λx:word16.match lt_w16 x addr with
+  (* e' prima di source: carica da old_mem *)
+  [ true ⇒ old_mem x
+  | false ⇒ match eq_w16 x addr with
+   (* la locazione corrisponde al punto corrente di source *)
+   [ true ⇒ hd
+   (* la locazione e' piu' avanti: ricorsione *)
+   | false ⇒ (mf_load_from_source_at old_mem tl (plus_w16nc addr 〈〈x0,x0〉:〈x0,x1〉〉)) x
+   ]
+  ]
+ ].
+
+(* ********************** *)
+(* TEOREMI/LEMMMI/ASSIOMI *)
+(* ********************** *)
+
+(*
+lemma mem_update_mem_update_a_a:
+ ∀s,a,v1,v2,b.
+  mem_update (mem_update s a v1) a v2 b = mem_update s a v2 b.
+ intros;
+ unfold mem_update;
+ unfold mem_update;
+ elim (eqb b a);
+ reflexivity.
+qed.
+
+lemma mem_update_mem_update_a_b:
+ ∀s,a1,v1,a2,v2,b.
+  a1 ≠ a2 →
+   mem_update (mem_update s a1 v1) a2 v2 b = mem_update (mem_update s a2 v2) a1 v1 b.
+ intros;
+ unfold mem_update;
+ unfold mem_update;
+ apply (bool_elim ? (eqb b a1)); intros;
+ apply (bool_elim ? (eqb b a2)); intros;
+ simplify;
+ [ elim H;
+   rewrite < (eqb_true_to_eq ? ? H1);
+   apply eqb_true_to_eq;
+   assumption
+ |*: reflexivity
+ ].
+qed.
+
+lemma eq_update_s_a_sa: ∀s,a,b. update s a (s a) b = s b.
+ intros;
+ unfold update;
+ apply (bool_elim ? (eqb b a) ? ?); simplify; intros;
+  [ rewrite > (eqb_true_to_eq ? ? H);
+    reflexivity
+  | reflexivity
+  ]
+qed.
+
+lemma inj_update:
+ ∀s,s',a,v,b. (a ≠ b → s b = s' b) → update s a v b = update s' a v b.
+ intros;
+ unfold update;
+ apply (bool_elim ? (eqb b a) ? ?); simplify; intros;
+  [ reflexivity
+  | apply H;
+    intro;
+    autobatch
+  ]
+qed.
+
+lemma not_eq_a_b_to_eq_update_a_b: ∀s,a,b,v. a ≠ b → update s a v b = s b.
+ intros;
+ unfold update;
+ rewrite > not_eq_to_eqb_false; simplify;
+  [ reflexivity
+  | intro;
+    autobatch
+  ]
+qed.
+*)
diff --git a/helm/software/matita/library/freescale/memory_struct.ma b/helm/software/matita/library/freescale/memory_struct.ma
new file mode 100644 (file)
index 0000000..5c80b88
--- /dev/null
@@ -0,0 +1,653 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/memory_struct/".
+
+(*include "/media/VIRTUOSO/freescale/translation.ma".*)
+include "freescale/translation.ma".
+
+(* **************************** *)
+(* TIPI PER I MODULI DI MEMORIA *)
+(* **************************** *)
+
+(* tipi di memoria:RAM/ROM/non installata *)
+inductive memory_type : Type ≝
+  MEM_READ_ONLY: memory_type
+| MEM_READ_WRITE: memory_type
+| MEM_OUT_OF_BOUND: memory_type.
+
+(* **************** *)
+(* TIPO ARRAY DA 16 *)
+(* **************** *)
+
+(* definizione di un array omogeneo di dimensione 16 *)
+inductive Prod16T (T:Type) : Type ≝
+array_16T : T → T → T → T → T → T → T → T →
+            T → T → T → T → T → T → T → T →
+            Prod16T T.
+
+(* abbiamo gia' gli esadecimali come tipo induttivo quindi: *)
+(* posso definire un getter a matrice sull'array *)
+definition getn_array16T ≝
+λn:exadecim.λT:Type.λp:Prod16T T.
+ match p with 
+  [ array_16T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15 ⇒
+ match n with
+  [ x0 ⇒ e00 | x1 ⇒ e01 | x2 ⇒ e02 | x3 ⇒ e03 | x4 ⇒ e04 | x5 ⇒ e05 | x6 ⇒ e06 | x7 ⇒ e07
+  | x8 ⇒ e08 | x9 ⇒ e09 | xA ⇒ e10 | xB ⇒ e11 | xC ⇒ e12 | xD ⇒ e13 | xE ⇒ e14 | xF ⇒ e15
+  ]].
+
+(* abbiamo gia' gli esadecimali come tipo induttivo quindi: *)
+(* posso definire un setter a matrice sull'array *)
+definition setn_array16T ≝
+λn:exadecim.λT:Type.λp:Prod16T T.λv:T.
+ match p with 
+  [ array_16T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15 ⇒
+ match n with
+  [ x0 ⇒ array_16T T v   e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+  | x1 ⇒ array_16T T e00 v   e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+  | x2 ⇒ array_16T T e00 e01 v   e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+  | x3 ⇒ array_16T T e00 e01 e02 v   e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+  | x4 ⇒ array_16T T e00 e01 e02 e03 v   e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+  | x5 ⇒ array_16T T e00 e01 e02 e03 e04 v   e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+  | x6 ⇒ array_16T T e00 e01 e02 e03 e04 e05 v   e07 e08 e09 e10 e11 e12 e13 e14 e15
+  | x7 ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 v   e08 e09 e10 e11 e12 e13 e14 e15
+  | x8 ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 v   e09 e10 e11 e12 e13 e14 e15
+  | x9 ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 v   e10 e11 e12 e13 e14 e15
+  | xA ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 v   e11 e12 e13 e14 e15
+  | xB ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 v   e12 e13 e14 e15
+  | xC ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 v   e13 e14 e15
+  | xD ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 v   e14 e15
+  | xE ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 v   e15
+  | xF ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 v
+  ]].
+
+(* abbiamo gia' gli esadecimali come tipo induttivo quindi: *)
+(* posso definire un setter multiplo [m,n] a matrice sull'array *)
+definition setmn_array16T ≝
+λm,n:exadecim.λT:Type.λp:Prod16T T.λv:T.
+ match p with 
+  [ array_16T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15 ⇒
+ match m with
+  [ x0 ⇒ match n with
+   [ x0 ⇒ array_16T T v e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x1 ⇒ array_16T T v v   e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x2 ⇒ array_16T T v v   v   e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x3 ⇒ array_16T T v v   v   v   e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x4 ⇒ array_16T T v v   v   v   v   e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x5 ⇒ array_16T T v v   v   v   v   v   e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x6 ⇒ array_16T T v v   v   v   v   v   v   e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x7 ⇒ array_16T T v v   v   v   v   v   v   v   e08 e09 e10 e11 e12 e13 e14 e15
+   | x8 ⇒ array_16T T v v   v   v   v   v   v   v   v   e09 e10 e11 e12 e13 e14 e15
+   | x9 ⇒ array_16T T v v   v   v   v   v   v   v   v   v   e10 e11 e12 e13 e14 e15
+   | xA ⇒ array_16T T v v   v   v   v   v   v   v   v   v   v   e11 e12 e13 e14 e15
+   | xB ⇒ array_16T T v v   v   v   v   v   v   v   v   v   v   v   e12 e13 e14 e15
+   | xC ⇒ array_16T T v v   v   v   v   v   v   v   v   v   v   v   v   e13 e14 e15
+   | xD ⇒ array_16T T v v   v   v   v   v   v   v   v   v   v   v   v   v   e14 e15
+   | xE ⇒ array_16T T v v   v   v   v   v   v   v   v   v   v   v   v   v   v   e15
+   | xF ⇒ array_16T T v v   v   v   v   v   v   v   v   v   v   v   v   v   v   v ]
+  | x1 ⇒ match n with
+   [ x0 ⇒ p
+   | x1 ⇒ array_16T T e00 v e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x2 ⇒ array_16T T e00 v v   e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x3 ⇒ array_16T T e00 v v   v   e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x4 ⇒ array_16T T e00 v v   v   v   e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x5 ⇒ array_16T T e00 v v   v   v   v   e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x6 ⇒ array_16T T e00 v v   v   v   v   v   e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x7 ⇒ array_16T T e00 v v   v   v   v   v   v   e08 e09 e10 e11 e12 e13 e14 e15
+   | x8 ⇒ array_16T T e00 v v   v   v   v   v   v   v   e09 e10 e11 e12 e13 e14 e15
+   | x9 ⇒ array_16T T e00 v v   v   v   v   v   v   v   v   e10 e11 e12 e13 e14 e15
+   | xA ⇒ array_16T T e00 v v   v   v   v   v   v   v   v   v   e11 e12 e13 e14 e15
+   | xB ⇒ array_16T T e00 v v   v   v   v   v   v   v   v   v   v   e12 e13 e14 e15
+   | xC ⇒ array_16T T e00 v v   v   v   v   v   v   v   v   v   v   v   e13 e14 e15
+   | xD ⇒ array_16T T e00 v v   v   v   v   v   v   v   v   v   v   v   v   e14 e15
+   | xE ⇒ array_16T T e00 v v   v   v   v   v   v   v   v   v   v   v   v   v   e15
+   | xF ⇒ array_16T T e00 v v   v   v   v   v   v   v   v   v   v   v   v   v   v ]
+  | x2 ⇒ match n with
+   [ x0 ⇒ p | x1 ⇒ p
+   | x2 ⇒ array_16T T e00 e01 v e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x3 ⇒ array_16T T e00 e01 v v   e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x4 ⇒ array_16T T e00 e01 v v   v   e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x5 ⇒ array_16T T e00 e01 v v   v   v   e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x6 ⇒ array_16T T e00 e01 v v   v   v   v   e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x7 ⇒ array_16T T e00 e01 v v   v   v   v   v   e08 e09 e10 e11 e12 e13 e14 e15
+   | x8 ⇒ array_16T T e00 e01 v v   v   v   v   v   v   e09 e10 e11 e12 e13 e14 e15
+   | x9 ⇒ array_16T T e00 e01 v v   v   v   v   v   v   v   e10 e11 e12 e13 e14 e15
+   | xA ⇒ array_16T T e00 e01 v v   v   v   v   v   v   v   v   e11 e12 e13 e14 e15
+   | xB ⇒ array_16T T e00 e01 v v   v   v   v   v   v   v   v   v   e12 e13 e14 e15
+   | xC ⇒ array_16T T e00 e01 v v   v   v   v   v   v   v   v   v   v   e13 e14 e15
+   | xD ⇒ array_16T T e00 e01 v v   v   v   v   v   v   v   v   v   v   v   e14 e15
+   | xE ⇒ array_16T T e00 e01 v v   v   v   v   v   v   v   v   v   v   v   v   e15
+   | xF ⇒ array_16T T e00 e01 v v   v   v   v   v   v   v   v   v   v   v   v   v ]
+  | x3 ⇒ match n with
+   [ x0 ⇒ p | x1 ⇒ p | x2 ⇒ p
+   | x3 ⇒ array_16T T e00 e01 e02 v e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x4 ⇒ array_16T T e00 e01 e02 v v   e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x5 ⇒ array_16T T e00 e01 e02 v v   v   e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x6 ⇒ array_16T T e00 e01 e02 v v   v   v   e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x7 ⇒ array_16T T e00 e01 e02 v v   v   v   v   e08 e09 e10 e11 e12 e13 e14 e15
+   | x8 ⇒ array_16T T e00 e01 e02 v v   v   v   v   v   e09 e10 e11 e12 e13 e14 e15
+   | x9 ⇒ array_16T T e00 e01 e02 v v   v   v   v   v   v   e10 e11 e12 e13 e14 e15
+   | xA ⇒ array_16T T e00 e01 e02 v v   v   v   v   v   v   v   e11 e12 e13 e14 e15
+   | xB ⇒ array_16T T e00 e01 e02 v v   v   v   v   v   v   v   v   e12 e13 e14 e15
+   | xC ⇒ array_16T T e00 e01 e02 v v   v   v   v   v   v   v   v   v   e13 e14 e15
+   | xD ⇒ array_16T T e00 e01 e02 v v   v   v   v   v   v   v   v   v   v   e14 e15
+   | xE ⇒ array_16T T e00 e01 e02 v v   v   v   v   v   v   v   v   v   v   v   e15
+   | xF ⇒ array_16T T e00 e01 e02 v v   v   v   v   v   v   v   v   v   v   v   v ]
+  | x4 ⇒ match n with
+   [ x0 ⇒ p | x1 ⇒ p | x2 ⇒ p | x3 ⇒ p
+   | x4 ⇒ array_16T T e00 e01 e02 e03 v e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x5 ⇒ array_16T T e00 e01 e02 e03 v v   e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x6 ⇒ array_16T T e00 e01 e02 e03 v v   v   e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x7 ⇒ array_16T T e00 e01 e02 e03 v v   v   v   e08 e09 e10 e11 e12 e13 e14 e15
+   | x8 ⇒ array_16T T e00 e01 e02 e03 v v   v   v   v   e09 e10 e11 e12 e13 e14 e15
+   | x9 ⇒ array_16T T e00 e01 e02 e03 v v   v   v   v   v   e10 e11 e12 e13 e14 e15
+   | xA ⇒ array_16T T e00 e01 e02 e03 v v   v   v   v   v   v   e11 e12 e13 e14 e15
+   | xB ⇒ array_16T T e00 e01 e02 e03 v v   v   v   v   v   v   v   e12 e13 e14 e15
+   | xC ⇒ array_16T T e00 e01 e02 e03 v v   v   v   v   v   v   v   v   e13 e14 e15
+   | xD ⇒ array_16T T e00 e01 e02 e03 v v   v   v   v   v   v   v   v   v   e14 e15
+   | xE ⇒ array_16T T e00 e01 e02 e03 v v   v   v   v   v   v   v   v   v   v   e15
+   | xF ⇒ array_16T T e00 e01 e02 e03 v v   v   v   v   v   v   v   v   v   v   v ]
+  | x5 ⇒ match n with
+   [ x0 ⇒ p | x1 ⇒ p | x2 ⇒ p | x3 ⇒ p | x4 ⇒ p
+   | x5 ⇒ array_16T T e00 e01 e02 e03 e04 v e06 e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x6 ⇒ array_16T T e00 e01 e02 e03 e04 v v   e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x7 ⇒ array_16T T e00 e01 e02 e03 e04 v v   v   e08 e09 e10 e11 e12 e13 e14 e15
+   | x8 ⇒ array_16T T e00 e01 e02 e03 e04 v v   v   v   e09 e10 e11 e12 e13 e14 e15
+   | x9 ⇒ array_16T T e00 e01 e02 e03 e04 v v   v   v   v   e10 e11 e12 e13 e14 e15
+   | xA ⇒ array_16T T e00 e01 e02 e03 e04 v v   v   v   v   v   e11 e12 e13 e14 e15
+   | xB ⇒ array_16T T e00 e01 e02 e03 e04 v v   v   v   v   v   v   e12 e13 e14 e15
+   | xC ⇒ array_16T T e00 e01 e02 e03 e04 v v   v   v   v   v   v   v   e13 e14 e15
+   | xD ⇒ array_16T T e00 e01 e02 e03 e04 v v   v   v   v   v   v   v   v   e14 e15
+   | xE ⇒ array_16T T e00 e01 e02 e03 e04 v v   v   v   v   v   v   v   v   v   e15
+   | xF ⇒ array_16T T e00 e01 e02 e03 e04 v v   v   v   v   v   v   v   v   v   v ]
+  | x6 ⇒ match n with
+   [ x0 ⇒ p | x1 ⇒ p | x2 ⇒ p | x3 ⇒ p | x4 ⇒ p | x5 ⇒ p
+   | x6 ⇒ array_16T T e00 e01 e02 e03 e04 e05 v e07 e08 e09 e10 e11 e12 e13 e14 e15
+   | x7 ⇒ array_16T T e00 e01 e02 e03 e04 e05 v v   e08 e09 e10 e11 e12 e13 e14 e15
+   | x8 ⇒ array_16T T e00 e01 e02 e03 e04 e05 v v   v   e09 e10 e11 e12 e13 e14 e15
+   | x9 ⇒ array_16T T e00 e01 e02 e03 e04 e05 v v   v   v   e10 e11 e12 e13 e14 e15
+   | xA ⇒ array_16T T e00 e01 e02 e03 e04 e05 v v   v   v   v   e11 e12 e13 e14 e15
+   | xB ⇒ array_16T T e00 e01 e02 e03 e04 e05 v v   v   v   v   v   e12 e13 e14 e15
+   | xC ⇒ array_16T T e00 e01 e02 e03 e04 e05 v v   v   v   v   v   v   e13 e14 e15
+   | xD ⇒ array_16T T e00 e01 e02 e03 e04 e05 v v   v   v   v   v   v   v   e14 e15
+   | xE ⇒ array_16T T e00 e01 e02 e03 e04 e05 v v   v   v   v   v   v   v   v   e15
+   | xF ⇒ array_16T T e00 e01 e02 e03 e04 e05 v v   v   v   v   v   v   v   v   v ]
+  | x7 ⇒ match n with
+   [ x0 ⇒ p | x1 ⇒ p | x2 ⇒ p | x3 ⇒ p | x4 ⇒ p | x5 ⇒ p | x6 ⇒ p
+   | x7 ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 v e08 e09 e10 e11 e12 e13 e14 e15
+   | x8 ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 v v   e09 e10 e11 e12 e13 e14 e15
+   | x9 ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 v v   v   e10 e11 e12 e13 e14 e15
+   | xA ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 v v   v   v   e11 e12 e13 e14 e15
+   | xB ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 v v   v   v   v   e12 e13 e14 e15
+   | xC ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 v v   v   v   v   v   e13 e14 e15
+   | xD ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 v v   v   v   v   v   v   e14 e15
+   | xE ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 v v   v   v   v   v   v   v   e15
+   | xF ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 v v   v   v   v   v   v   v   v ]
+  | x8 ⇒ match n with
+   [ x0 ⇒ p | x1 ⇒ p | x2 ⇒ p | x3 ⇒ p | x4 ⇒ p | x5 ⇒ p | x6 ⇒ p | x7 ⇒ p
+   | x8 ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 v e09 e10 e11 e12 e13 e14 e15
+   | x9 ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 v v   e10 e11 e12 e13 e14 e15
+   | xA ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 v v   v   e11 e12 e13 e14 e15
+   | xB ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 v v   v   v   e12 e13 e14 e15
+   | xC ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 v v   v   v   v   e13 e14 e15
+   | xD ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 v v   v   v   v   v   e14 e15
+   | xE ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 v v   v   v   v   v   v   e15
+   | xF ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 v v   v   v   v   v   v   v ]
+  | x9 ⇒ match n with
+   [ x0 ⇒ p | x1 ⇒ p | x2 ⇒ p  | x3 ⇒ p | x4 ⇒ p | x5 ⇒ p | x6 ⇒ p | x7 ⇒ p
+   | x8 ⇒ p
+   | x9 ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 v e10 e11 e12 e13 e14 e15
+   | xA ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 v v   e11 e12 e13 e14 e15
+   | xB ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 v v   v   e12 e13 e14 e15
+   | xC ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 v v   v   v   e13 e14 e15
+   | xD ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 v v   v   v   v   e14 e15
+   | xE ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 v v   v   v   v   v   e15
+   | xF ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 v v   v   v   v   v   v ]
+  | xA ⇒ match n with
+   [ x0 ⇒ p | x1 ⇒ p | x2 ⇒ p | x3 ⇒ p | x4 ⇒ p | x5 ⇒ p | x6 ⇒ p | x7 ⇒ p
+   | x8 ⇒ p | x9 ⇒ p
+   | xA ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 v e11 e12 e13 e14 e15
+   | xB ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 v v   e12 e13 e14 e15
+   | xC ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 v v   v   e13 e14 e15
+   | xD ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 v v   v   v   e14 e15
+   | xE ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 v v   v   v   v   e15
+   | xF ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 v v   v   v   v   v ]
+  | xB ⇒ match n with
+   [ x0 ⇒ p | x1 ⇒ p | x2 ⇒ p | x3 ⇒ p | x4 ⇒ p | x5 ⇒ p | x6 ⇒ p | x7 ⇒ p
+   | x8 ⇒ p | x9 ⇒ p | xA ⇒ p
+   | xB ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 v e12 e13 e14 e15
+   | xC ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 v v   e13 e14 e15
+   | xD ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 v v   v   e14 e15
+   | xE ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 v v   v   v   e15
+   | xF ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 v v   v   v   v ]
+  | xC ⇒ match n with
+   [ x0 ⇒ p | x1 ⇒ p | x2 ⇒ p | x3 ⇒ p | x4 ⇒ p | x5 ⇒ p | x6 ⇒ p | x7 ⇒ p
+   | x8 ⇒ p | x9 ⇒ p | xA ⇒ p | xB ⇒ p
+   | xC ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 v e13 e14 e15
+   | xD ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 v v   e14 e15
+   | xE ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 v v   v   e15
+   | xF ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 v v   v   v ]
+  | xD ⇒ match n with
+   [ x0 ⇒ p | x1 ⇒ p | x2 ⇒ p | x3 ⇒ p | x4 ⇒ p | x5 ⇒ p | x6 ⇒ p | x7 ⇒ p
+   | x8 ⇒ p | x9 ⇒ p | xA ⇒ p | xB ⇒ p | xC ⇒ p
+   | xD ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 v e14 e15
+   | xE ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 v v   e15
+   | xF ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 v v   v ]
+  | xE ⇒ match n with
+   [ x0 ⇒ p | x1 ⇒ p | x2 ⇒ p | x3 ⇒ p | x4 ⇒ p | x5 ⇒ p | x6 ⇒ p | x7 ⇒ p
+   | x8 ⇒ p | x9 ⇒ p | xA ⇒ p | xB ⇒ p | xC ⇒ p | xD ⇒ p
+   | xE ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 v e15
+   | xF ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 v v ]
+  | xF ⇒ match n with
+   [ x0 ⇒ p | x1 ⇒ p | x2 ⇒ p | x3 ⇒ p | x4 ⇒ p | x5 ⇒ p | x6 ⇒ p | x7 ⇒ p
+   | x8 ⇒ p | x9 ⇒ p | xA ⇒ p | xB ⇒ p | xC ⇒ p | xD ⇒ p | xE ⇒ p
+   | xF ⇒ array_16T T e00 e01 e02 e03 e04 e05 e06 e07 e08 e09 e10 e11 e12 e13 e14 v ]
+  ]].
+
+(* abbiamo gia' gli esadecimali come tipo induttivo quindi: *)
+(* posso definire un setter composto [m+1,n-1] a matrice sull'array *)
+(* NB: obbiettivo evitare l'overflow *)
+definition setmn_array16T_succ_pred ≝
+λm,n:exadecim.λT:Type.λp:Prod16T T.λv:T.
+ match lt_ex m xF with
+  [ true ⇒ match gt_ex n x0 with
+   [ true ⇒ setmn_array16T (succ_ex m) (pred_ex n) T p v
+   | false ⇒ p
+   ]
+  | false ⇒ p
+  ].
+
+(* abbiamo gia' gli esadecimali come tipo induttivo quindi: *)
+(* posso definire un setter composto [m+1,F] a matrice sull'array *)
+(* NB: obbiettivo evitare l'overflow *)
+definition setmn_array16T_succ ≝
+λm:exadecim.λT:Type.λp:Prod16T T.λv:T.
+ match lt_ex m xF with
+  [ true ⇒ setmn_array16T (succ_ex m) xF T p v
+  | false ⇒ p
+  ].
+
+(* abbiamo gia' gli esadecimali come tipo induttivo quindi: *)
+(* posso definire un setter composto [0,n-1] a matrice sull'array *)
+(* NB: obbiettivo evitare l'overflow *)
+definition setmn_array16T_pred ≝
+λn:exadecim.λT:Type.λp:Prod16T T.λv:T.
+ match gt_ex n x0 with
+  [ true ⇒ setmn_array16T x0 (pred_ex n) T p v
+  | false ⇒ p
+  ].
+
+(* ************************** *)
+(* TIPO BYTE COME INSIEME BIT *)
+(* ************************** *)
+
+(* definizione di un byte come 8 bit *)
+inductive Prod8T (T:Type) : Type ≝
+array_8T : T → T → T → T → T → T → T → T →
+           Prod8T T.
+
+(* abbiamo gia' gli ottali come tipo induttivo quindi: *)
+(* posso definire un getter a matrice sull'array *)
+definition getn_array8T ≝
+λn:oct.λT:Type.λp:Prod8T T.
+ match p with 
+  [ array_8T e07 e06 e05 e04 e03 e02 e01 e00 ⇒
+ match n with
+  [ o0 ⇒ e00 | o1 ⇒ e01 | o2 ⇒ e02 | o3 ⇒ e03 | o4 ⇒ e04 | o5 ⇒ e05 | o6 ⇒ e06 | o7 ⇒ e07 ]].
+
+(* abbiamo gia' gli ottali come tipo induttivo quindi: *)
+(* posso definire un setter a matrice sull'array *)
+definition setn_array8T ≝
+λn:oct.λT:Type.λp:Prod8T T.λv:T.
+ match p with 
+  [ array_8T e07 e06 e05 e04 e03 e02 e01 e00 ⇒
+ match n with
+  [ o0 ⇒ array_8T T e07 e06 e05 e04 e03 e02 e01 v
+  | o1 ⇒ array_8T T e07 e06 e05 e04 e03 e02 v   e00
+  | o2 ⇒ array_8T T e07 e06 e05 e04 e03 v   e01 e00
+  | o3 ⇒ array_8T T e07 e06 e05 e04 v   e02 e01 e00
+  | o4 ⇒ array_8T T e07 e06 e05 v   e03 e02 e01 e00
+  | o5 ⇒ array_8T T e07 e06 v   e04 e03 e02 e01 e00
+  | o6 ⇒ array_8T T e07 v   e05 e04 e03 e02 e01 e00
+  | o7 ⇒ array_8T T v   e06 e05 e04 e03 e02 e01 e00
+  ]].
+
+(* lettura byte *)
+definition byte8_of_bits ≝
+λp:Prod8T bool.
+ match p with 
+  [ array_8T e07 e06 e05 e04 e03 e02 e01 e00 ⇒
+   mk_byte8
+   (or_ex (match e07 with [ true ⇒ x8 | false ⇒ x0 ])
+   (or_ex (match e06 with [ true ⇒ x4 | false ⇒ x0 ])
+   (or_ex (match e05 with [ true ⇒ x2 | false ⇒ x0 ])
+          (match e04 with [ true ⇒ x1 | false ⇒ x0 ]))))
+   (or_ex (match e03 with [ true ⇒ x8 | false ⇒ x0 ])
+   (or_ex (match e02 with [ true ⇒ x4 | false ⇒ x0 ])
+   (or_ex (match e01 with [ true ⇒ x2 | false ⇒ x0 ])
+          (match e00 with [ true ⇒ x1 | false ⇒ x0 ])))) ].
+
+(* scrittura byte *)
+definition bits_of_byte8 ≝
+λp:byte8.
+ match b8h p with
+  [ x0 ⇒ match b8l p with
+   [ x0 ⇒ array_8T bool false false false false false false false false
+   | x1 ⇒ array_8T bool false false false false false false false true
+   | x2 ⇒ array_8T bool false false false false false false true  false
+   | x3 ⇒ array_8T bool false false false false false false true  true
+   | x4 ⇒ array_8T bool false false false false false true  false false
+   | x5 ⇒ array_8T bool false false false false false true  false true
+   | x6 ⇒ array_8T bool false false false false false true  true  false
+   | x7 ⇒ array_8T bool false false false false false true  true  true
+   | x8 ⇒ array_8T bool false false false false true  false false false
+   | x9 ⇒ array_8T bool false false false false true  false false true
+   | xA ⇒ array_8T bool false false false false true  false true  false
+   | xB ⇒ array_8T bool false false false false true  false true  true
+   | xC ⇒ array_8T bool false false false false true  true  false false
+   | xD ⇒ array_8T bool false false false false true  true  false true
+   | xE ⇒ array_8T bool false false false false true  true  true  false
+   | xF ⇒ array_8T bool false false false false true  true  true  true ]
+  | x1 ⇒ match b8l p with
+   [ x0 ⇒ array_8T bool false false false true  false false false false
+   | x1 ⇒ array_8T bool false false false true  false false false true
+   | x2 ⇒ array_8T bool false false false true  false false true  false
+   | x3 ⇒ array_8T bool false false false true  false false true  true
+   | x4 ⇒ array_8T bool false false false true  false true  false false
+   | x5 ⇒ array_8T bool false false false true  false true  false true
+   | x6 ⇒ array_8T bool false false false true  false true  true  false
+   | x7 ⇒ array_8T bool false false false true  false true  true  true
+   | x8 ⇒ array_8T bool false false false true  true  false false false
+   | x9 ⇒ array_8T bool false false false true  true  false false true
+   | xA ⇒ array_8T bool false false false true  true  false true  false
+   | xB ⇒ array_8T bool false false false true  true  false true  true
+   | xC ⇒ array_8T bool false false false true  true  true  false false
+   | xD ⇒ array_8T bool false false false true  true  true  false true
+   | xE ⇒ array_8T bool false false false true  true  true  true  false
+   | xF ⇒ array_8T bool false false false true  true  true  true  true ]
+  | x2 ⇒ match b8l p with
+   [ x0 ⇒ array_8T bool false false true  false false false false false
+   | x1 ⇒ array_8T bool false false true  false false false false true
+   | x2 ⇒ array_8T bool false false true  false false false true  false
+   | x3 ⇒ array_8T bool false false true  false false false true  true
+   | x4 ⇒ array_8T bool false false true  false false true  false false
+   | x5 ⇒ array_8T bool false false true  false false true  false true
+   | x6 ⇒ array_8T bool false false true  false false true  true  false
+   | x7 ⇒ array_8T bool false false true  false false true  true  true
+   | x8 ⇒ array_8T bool false false true  false true  false false false
+   | x9 ⇒ array_8T bool false false true  false true  false false true
+   | xA ⇒ array_8T bool false false true  false true  false true  false
+   | xB ⇒ array_8T bool false false true  false true  false true  true
+   | xC ⇒ array_8T bool false false true  false true  true  false false
+   | xD ⇒ array_8T bool false false true  false true  true  false true
+   | xE ⇒ array_8T bool false false true  false true  true  true  false
+   | xF ⇒ array_8T bool false false true  false true  true  true  true ]
+  | x3 ⇒ match b8l p with
+   [ x0 ⇒ array_8T bool false false true  true  false false false false
+   | x1 ⇒ array_8T bool false false true  true  false false false true
+   | x2 ⇒ array_8T bool false false true  true  false false true  false
+   | x3 ⇒ array_8T bool false false true  true  false false true  true
+   | x4 ⇒ array_8T bool false false true  true  false true  false false
+   | x5 ⇒ array_8T bool false false true  true  false true  false true
+   | x6 ⇒ array_8T bool false false true  true  false true  true  false
+   | x7 ⇒ array_8T bool false false true  true  false true  true  true
+   | x8 ⇒ array_8T bool false false true  true  true  false false false
+   | x9 ⇒ array_8T bool false false true  true  true  false false true
+   | xA ⇒ array_8T bool false false true  true  true  false true  false
+   | xB ⇒ array_8T bool false false true  true  true  false true  true
+   | xC ⇒ array_8T bool false false true  true  true  true  false false
+   | xD ⇒ array_8T bool false false true  true  true  true  false true
+   | xE ⇒ array_8T bool false false true  true  true  true  true  false
+   | xF ⇒ array_8T bool false false true  true  true  true  true  true ]
+  | x4 ⇒ match b8l p with
+   [ x0 ⇒ array_8T bool false true  false false false false false false
+   | x1 ⇒ array_8T bool false true  false false false false false true
+   | x2 ⇒ array_8T bool false true  false false false false true  false
+   | x3 ⇒ array_8T bool false true  false false false false true  true
+   | x4 ⇒ array_8T bool false true  false false false true  false false
+   | x5 ⇒ array_8T bool false true  false false false true  false true
+   | x6 ⇒ array_8T bool false true  false false false true  true  false
+   | x7 ⇒ array_8T bool false true  false false false true  true  true
+   | x8 ⇒ array_8T bool false true  false false true  false false false
+   | x9 ⇒ array_8T bool false true  false false true  false false true
+   | xA ⇒ array_8T bool false true  false false true  false true  false
+   | xB ⇒ array_8T bool false true  false false true  false true  true
+   | xC ⇒ array_8T bool false true  false false true  true  false false
+   | xD ⇒ array_8T bool false true  false false true  true  false true
+   | xE ⇒ array_8T bool false true  false false true  true  true  false
+   | xF ⇒ array_8T bool false true  false false true  true  true  true ]
+  | x5 ⇒ match b8l p with
+   [ x0 ⇒ array_8T bool false true  false true  false false false false
+   | x1 ⇒ array_8T bool false true  false true  false false false true
+   | x2 ⇒ array_8T bool false true  false true  false false true  false
+   | x3 ⇒ array_8T bool false true  false true  false false true  true
+   | x4 ⇒ array_8T bool false true  false true  false true  false false
+   | x5 ⇒ array_8T bool false true  false true  false true  false true
+   | x6 ⇒ array_8T bool false true  false true  false true  true  false
+   | x7 ⇒ array_8T bool false true  false true  false true  true  true
+   | x8 ⇒ array_8T bool false true  false true  true  false false false
+   | x9 ⇒ array_8T bool false true  false true  true  false false true
+   | xA ⇒ array_8T bool false true  false true  true  false true  false
+   | xB ⇒ array_8T bool false true  false true  true  false true  true
+   | xC ⇒ array_8T bool false true  false true  true  true  false false
+   | xD ⇒ array_8T bool false true  false true  true  true  false true
+   | xE ⇒ array_8T bool false true  false true  true  true  true  false
+   | xF ⇒ array_8T bool false true  false true  true  true  true  true ]
+  | x6 ⇒ match b8l p with
+   [ x0 ⇒ array_8T bool false true  true  false false false false false
+   | x1 ⇒ array_8T bool false true  true  false false false false true
+   | x2 ⇒ array_8T bool false true  true  false false false true  false
+   | x3 ⇒ array_8T bool false true  true  false false false true  true
+   | x4 ⇒ array_8T bool false true  true  false false true  false false
+   | x5 ⇒ array_8T bool false true  true  false false true  false true
+   | x6 ⇒ array_8T bool false true  true  false false true  true  false
+   | x7 ⇒ array_8T bool false true  true  false false true  true  true
+   | x8 ⇒ array_8T bool false true  true  false true  false false false
+   | x9 ⇒ array_8T bool false true  true  false true  false false true
+   | xA ⇒ array_8T bool false true  true  false true  false true  false
+   | xB ⇒ array_8T bool false true  true  false true  false true  true
+   | xC ⇒ array_8T bool false true  true  false true  true  false false
+   | xD ⇒ array_8T bool false true  true  false true  true  false true
+   | xE ⇒ array_8T bool false true  true  false true  true  true  false
+   | xF ⇒ array_8T bool false true  true  false true  true  true  true ]
+  | x7 ⇒ match b8l p with
+   [ x0 ⇒ array_8T bool false true  true  true  false false false false
+   | x1 ⇒ array_8T bool false true  true  true  false false false true
+   | x2 ⇒ array_8T bool false true  true  true  false false true  false
+   | x3 ⇒ array_8T bool false true  true  true  false false true  true
+   | x4 ⇒ array_8T bool false true  true  true  false true  false false
+   | x5 ⇒ array_8T bool false true  true  true  false true  false true
+   | x6 ⇒ array_8T bool false true  true  true  false true  true  false
+   | x7 ⇒ array_8T bool false true  true  true  false true  true  true
+   | x8 ⇒ array_8T bool false true  true  true  true  false false false
+   | x9 ⇒ array_8T bool false true  true  true  true  false false true
+   | xA ⇒ array_8T bool false true  true  true  true  false true  false
+   | xB ⇒ array_8T bool false true  true  true  true  false true  true
+   | xC ⇒ array_8T bool false true  true  true  true  true  false false
+   | xD ⇒ array_8T bool false true  true  true  true  true  false true
+   | xE ⇒ array_8T bool false true  true  true  true  true  true  false
+   | xF ⇒ array_8T bool false true  true  true  true  true  true  true ]
+  | x8 ⇒ match b8l p with
+   [ x0 ⇒ array_8T bool true  false false false false false false false
+   | x1 ⇒ array_8T bool true  false false false false false false true
+   | x2 ⇒ array_8T bool true  false false false false false true  false
+   | x3 ⇒ array_8T bool true  false false false false false true  true
+   | x4 ⇒ array_8T bool true  false false false false true  false false
+   | x5 ⇒ array_8T bool true  false false false false true  false true
+   | x6 ⇒ array_8T bool true  false false false false true  true  false
+   | x7 ⇒ array_8T bool true  false false false false true  true  true
+   | x8 ⇒ array_8T bool true  false false false true  false false false
+   | x9 ⇒ array_8T bool true  false false false true  false false true
+   | xA ⇒ array_8T bool true  false false false true  false true  false
+   | xB ⇒ array_8T bool true  false false false true  false true  true
+   | xC ⇒ array_8T bool true  false false false true  true  false false
+   | xD ⇒ array_8T bool true  false false false true  true  false true
+   | xE ⇒ array_8T bool true  false false false true  true  true  false
+   | xF ⇒ array_8T bool true  false false false true  true  true  true ]
+  | x9 ⇒ match b8l p with
+   [ x0 ⇒ array_8T bool true  false false true  false false false false
+   | x1 ⇒ array_8T bool true  false false true  false false false true
+   | x2 ⇒ array_8T bool true  false false true  false false true  false
+   | x3 ⇒ array_8T bool true  false false true  false false true  true
+   | x4 ⇒ array_8T bool true  false false true  false true  false false
+   | x5 ⇒ array_8T bool true  false false true  false true  false true
+   | x6 ⇒ array_8T bool true  false false true  false true  true  false
+   | x7 ⇒ array_8T bool true  false false true  false true  true  true
+   | x8 ⇒ array_8T bool true  false false true  true  false false false
+   | x9 ⇒ array_8T bool true  false false true  true  false false true
+   | xA ⇒ array_8T bool true  false false true  true  false true  false
+   | xB ⇒ array_8T bool true  false false true  true  false true  true
+   | xC ⇒ array_8T bool true  false false true  true  true  false false
+   | xD ⇒ array_8T bool true  false false true  true  true  false true
+   | xE ⇒ array_8T bool true  false false true  true  true  true  false
+   | xF ⇒ array_8T bool true  false false true  true  true  true  true ]
+  | xA ⇒ match b8l p with
+   [ x0 ⇒ array_8T bool true  false true  false false false false false
+   | x1 ⇒ array_8T bool true  false true  false false false false true
+   | x2 ⇒ array_8T bool true  false true  false false false true  false
+   | x3 ⇒ array_8T bool true  false true  false false false true  true
+   | x4 ⇒ array_8T bool true  false true  false false true  false false
+   | x5 ⇒ array_8T bool true  false true  false false true  false true
+   | x6 ⇒ array_8T bool true  false true  false false true  true  false
+   | x7 ⇒ array_8T bool true  false true  false false true  true  true
+   | x8 ⇒ array_8T bool true  false true  false true  false false false
+   | x9 ⇒ array_8T bool true  false true  false true  false false true
+   | xA ⇒ array_8T bool true  false true  false true  false true  false
+   | xB ⇒ array_8T bool true  false true  false true  false true  true
+   | xC ⇒ array_8T bool true  false true  false true  true  false false
+   | xD ⇒ array_8T bool true  false true  false true  true  false true
+   | xE ⇒ array_8T bool true  false true  false true  true  true  false
+   | xF ⇒ array_8T bool true  false true  false true  true  true  true ]
+  | xB ⇒ match b8l p with
+   [ x0 ⇒ array_8T bool true  false true  true  false false false false
+   | x1 ⇒ array_8T bool true  false true  true  false false false true
+   | x2 ⇒ array_8T bool true  false true  true  false false true  false
+   | x3 ⇒ array_8T bool true  false true  true  false false true  true
+   | x4 ⇒ array_8T bool true  false true  true  false true  false false
+   | x5 ⇒ array_8T bool true  false true  true  false true  false true
+   | x6 ⇒ array_8T bool true  false true  true  false true  true  false
+   | x7 ⇒ array_8T bool true  false true  true  false true  true  true
+   | x8 ⇒ array_8T bool true  false true  true  true  false false false
+   | x9 ⇒ array_8T bool true  false true  true  true  false false true
+   | xA ⇒ array_8T bool true  false true  true  true  false true  false
+   | xB ⇒ array_8T bool true  false true  true  true  false true  true
+   | xC ⇒ array_8T bool true  false true  true  true  true  false false
+   | xD ⇒ array_8T bool true  false true  true  true  true  false true
+   | xE ⇒ array_8T bool true  false true  true  true  true  true  false
+   | xF ⇒ array_8T bool true  false true  true  true  true  true  true ]
+  | xC ⇒ match b8l p with
+   [ x0 ⇒ array_8T bool true  true  false false false false false false
+   | x1 ⇒ array_8T bool true  true  false false false false false true
+   | x2 ⇒ array_8T bool true  true  false false false false true  false
+   | x3 ⇒ array_8T bool true  true  false false false false true  true
+   | x4 ⇒ array_8T bool true  true  false false false true  false false
+   | x5 ⇒ array_8T bool true  true  false false false true  false true
+   | x6 ⇒ array_8T bool true  true  false false false true  true  false
+   | x7 ⇒ array_8T bool true  true  false false false true  true  true
+   | x8 ⇒ array_8T bool true  true  false false true  false false false
+   | x9 ⇒ array_8T bool true  true  false false true  false false true
+   | xA ⇒ array_8T bool true  true  false false true  false true  false
+   | xB ⇒ array_8T bool true  true  false false true  false true  true
+   | xC ⇒ array_8T bool true  true  false false true  true  false false
+   | xD ⇒ array_8T bool true  true  false false true  true  false true
+   | xE ⇒ array_8T bool true  true  false false true  true  true  false
+   | xF ⇒ array_8T bool true  true  false false true  true  true  true ]
+  | xD ⇒ match b8l p with
+   [ x0 ⇒ array_8T bool true  true  false true  false false false false
+   | x1 ⇒ array_8T bool true  true  false true  false false false true
+   | x2 ⇒ array_8T bool true  true  false true  false false true  false
+   | x3 ⇒ array_8T bool true  true  false true  false false true  true
+   | x4 ⇒ array_8T bool true  true  false true  false true  false false
+   | x5 ⇒ array_8T bool true  true  false true  false true  false true
+   | x6 ⇒ array_8T bool true  true  false true  false true  true  false
+   | x7 ⇒ array_8T bool true  true  false true  false true  true  true
+   | x8 ⇒ array_8T bool true  true  false true  true  false false false
+   | x9 ⇒ array_8T bool true  true  false true  true  false false true
+   | xA ⇒ array_8T bool true  true  false true  true  false true  false
+   | xB ⇒ array_8T bool true  true  false true  true  false true  true
+   | xC ⇒ array_8T bool true  true  false true  true  true  false false
+   | xD ⇒ array_8T bool true  true  false true  true  true  false true
+   | xE ⇒ array_8T bool true  true  false true  true  true  true  false
+   | xF ⇒ array_8T bool true  true  false true  true  true  true  true ]
+  | xE ⇒ match b8l p with
+   [ x0 ⇒ array_8T bool true  true  true  false false false false false
+   | x1 ⇒ array_8T bool true  true  true  false false false false true
+   | x2 ⇒ array_8T bool true  true  true  false false false true  false
+   | x3 ⇒ array_8T bool true  true  true  false false false true  true
+   | x4 ⇒ array_8T bool true  true  true  false false true  false false
+   | x5 ⇒ array_8T bool true  true  true  false false true  false true
+   | x6 ⇒ array_8T bool true  true  true  false false true  true  false
+   | x7 ⇒ array_8T bool true  true  true  false false true  true  true
+   | x8 ⇒ array_8T bool true  true  true  false true  false false false
+   | x9 ⇒ array_8T bool true  true  true  false true  false false true
+   | xA ⇒ array_8T bool true  true  true  false true  false true  false
+   | xB ⇒ array_8T bool true  true  true  false true  false true  true
+   | xC ⇒ array_8T bool true  true  true  false true  true  false false
+   | xD ⇒ array_8T bool true  true  true  false true  true  false true
+   | xE ⇒ array_8T bool true  true  true  false true  true  true  false
+   | xF ⇒ array_8T bool true  true  true  false true  true  true  true ]
+  | xF ⇒ match b8l p with
+   [ x0 ⇒ array_8T bool true  true  true  true  false false false false
+   | x1 ⇒ array_8T bool true  true  true  true  false false false true
+   | x2 ⇒ array_8T bool true  true  true  true  false false true  false
+   | x3 ⇒ array_8T bool true  true  true  true  false false true  true
+   | x4 ⇒ array_8T bool true  true  true  true  false true  false false
+   | x5 ⇒ array_8T bool true  true  true  true  false true  false true
+   | x6 ⇒ array_8T bool true  true  true  true  false true  true  false
+   | x7 ⇒ array_8T bool true  true  true  true  false true  true  true
+   | x8 ⇒ array_8T bool true  true  true  true  true  false false false
+   | x9 ⇒ array_8T bool true  true  true  true  true  false false true
+   | xA ⇒ array_8T bool true  true  true  true  true  false true  false
+   | xB ⇒ array_8T bool true  true  true  true  true  false true  true
+   | xC ⇒ array_8T bool true  true  true  true  true  true  false false
+   | xD ⇒ array_8T bool true  true  true  true  true  true  false true
+   | xE ⇒ array_8T bool true  true  true  true  true  true  true  false
+   | xF ⇒ array_8T bool true  true  true  true  true  true  true  true ]
+  ].
+
+(* ********************** *)
+(* TEOREMI/LEMMMI/ASSIOMI *)
+(* ********************** *)
+
+(* check *)
+lemma ok_byte8_of_bits_bits_of_byte8 :
+ ∀b.
+  byte8_of_bits (bits_of_byte8 b) = b.
+ intros;
+ elim b; elim e; elim e1;
+ reflexivity.
+qed.
+
+(* check *)
+lemma ok_bits_of_byte8_byte8_of_bits :
+ ∀b7,b6,b5,b4,b3,b2,b1,b0.
+  bits_of_byte8 (byte8_of_bits (array_8T bool b7 b6 b5 b4 b3 b2 b1 b0))
+  = array_8T bool b7 b6 b5 b4 b3 b2 b1 b0.
+ intros;
+ elim b0; elim b1; elim b2; elim b3; elim b4; elim b5; elim b6; elim b7;
+ reflexivity.
+qed.
diff --git a/helm/software/matita/library/freescale/memory_trees.ma b/helm/software/matita/library/freescale/memory_trees.ma
new file mode 100644 (file)
index 0000000..9fba644
--- /dev/null
@@ -0,0 +1,260 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/memory_trees/".
+
+(*include "/media/VIRTUOSO/freescale/memory_struct.ma".*)
+include "freescale/memory_struct.ma".
+
+(* ********************* *)
+(* MEMORIA E DESCRITTORE *)
+(* ********************* *)
+
+(* tutta la memoria non installata *)
+definition mt_out_of_bound_memory ≝
+let lev4 ≝ array_16T ? 
+           MEM_OUT_OF_BOUND MEM_OUT_OF_BOUND MEM_OUT_OF_BOUND MEM_OUT_OF_BOUND
+           MEM_OUT_OF_BOUND MEM_OUT_OF_BOUND MEM_OUT_OF_BOUND MEM_OUT_OF_BOUND
+           MEM_OUT_OF_BOUND MEM_OUT_OF_BOUND MEM_OUT_OF_BOUND MEM_OUT_OF_BOUND
+           MEM_OUT_OF_BOUND MEM_OUT_OF_BOUND MEM_OUT_OF_BOUND MEM_OUT_OF_BOUND
+           in
+let lev3 ≝ array_16T ?
+           lev4 lev4 lev4 lev4 lev4 lev4 lev4 lev4
+           lev4 lev4 lev4 lev4 lev4 lev4 lev4 lev4
+           in
+let lev2 ≝ array_16T ?
+           lev3 lev3 lev3 lev3 lev3 lev3 lev3 lev3
+           lev3 lev3 lev3 lev3 lev3 lev3 lev3 lev3
+           in
+let lev1 ≝ array_16T ?
+           lev2 lev2 lev2 lev2 lev2 lev2 lev2 lev2
+           lev2 lev2 lev2 lev2 lev2 lev2 lev2 lev2
+           in
+lev1.
+
+(* tutta la memoria a 0 *)
+definition mt_zero_memory ≝
+let lev4 ≝ array_16T ? 
+           (mk_byte8 x0 x0) (mk_byte8 x0 x0) (mk_byte8 x0 x0) (mk_byte8 x0 x0)
+           (mk_byte8 x0 x0) (mk_byte8 x0 x0) (mk_byte8 x0 x0) (mk_byte8 x0 x0)
+           (mk_byte8 x0 x0) (mk_byte8 x0 x0) (mk_byte8 x0 x0) (mk_byte8 x0 x0)
+           (mk_byte8 x0 x0) (mk_byte8 x0 x0) (mk_byte8 x0 x0) (mk_byte8 x0 x0)
+           in
+let lev3 ≝ array_16T ?
+           lev4 lev4 lev4 lev4 lev4 lev4 lev4 lev4
+           lev4 lev4 lev4 lev4 lev4 lev4 lev4 lev4
+           in
+let lev2 ≝ array_16T ?
+           lev3 lev3 lev3 lev3 lev3 lev3 lev3 lev3
+           lev3 lev3 lev3 lev3 lev3 lev3 lev3 lev3
+           in
+let lev1 ≝ array_16T ?
+           lev2 lev2 lev2 lev2 lev2 lev2 lev2 lev2
+           lev2 lev2 lev2 lev2 lev2 lev2 lev2 lev2
+           in
+lev1.
+
+(* visita di un albero da 64KB di elementi: ln16(65536)=4 passaggi *)
+definition mt_visit ≝
+λT:Type.λdata:Prod16T (Prod16T (Prod16T (Prod16T T))).λaddr:word16.
+ match addr with
+  [ mk_word16 wh wl ⇒
+ getn_array16T (b8l wl) ?
+  (getn_array16T (b8h wl) ?
+   (getn_array16T (b8l wh) ?
+    (getn_array16T (b8h wh) ? data))) ].
+
+(* scrittura di un elemento in un albero da 64KB *)
+definition mt_update ≝
+λT:Type.λdata:Prod16T (Prod16T (Prod16T (Prod16T T))).λaddr:word16.λv:T.
+ match addr with
+  [ mk_word16 wh wl ⇒
+ let lev2 ≝ getn_array16T (b8h wh) ? data in
+ let lev3 ≝ getn_array16T (b8l wh) ? lev2 in
+ let lev4 ≝ getn_array16T (b8h wl) ? lev3 in
+ setn_array16T (b8h wh) ? data
+  (setn_array16T (b8l wh) ? lev2
+   (setn_array16T (b8h wl) ? lev3
+    (setn_array16T (b8l wl) T lev4 v))) ].
+
+(* scrittura di un range in un albero da 64KB *)
+definition mt_update_ranged ≝
+λT:Type.λdata:Prod16T (Prod16T (Prod16T (Prod16T T))).λi,s:word16.λv:T.
+ (* ok i≤s *)
+ match le_w16 i s with
+  [ true ⇒
+ match i with
+  [ mk_word16 ih il ⇒
+ match s with
+  [ mk_word16 sh sl ⇒
+ let aux_4 ≝ Prod16T T in
+ let aux_3 ≝ Prod16T (Prod16T T) in
+ let aux_2 ≝ Prod16T (Prod16T (Prod16T T)) in
+
+ let ilev2 ≝ getn_array16T (b8h ih) aux_2 data in
+ let ilev3 ≝ getn_array16T (b8l ih) aux_3 ilev2 in
+ let ilev4 ≝ getn_array16T (b8h il) aux_4 ilev3 in
+
+ let slev2 ≝ getn_array16T (b8h sh) aux_2 data in
+ let slev3 ≝ getn_array16T (b8l sh) aux_3 slev2 in
+ let slev4 ≝ getn_array16T (b8h sl) aux_4 slev3 in
+
+ let vlev4 ≝ array_16T T v v v v v v v v v v v v v v v v in
+ let vlev3 ≝ array_16T aux_4 vlev4 vlev4 vlev4 vlev4 vlev4 vlev4 vlev4 vlev4
+                             vlev4 vlev4 vlev4 vlev4 vlev4 vlev4 vlev4 vlev4 in
+ let vlev2 ≝ array_16T aux_3 vlev3 vlev3 vlev3 vlev3 vlev3 vlev3 vlev3 vlev3
+                             vlev3 vlev3 vlev3 vlev3 vlev3 vlev3 vlev3 vlev3 in
+
+ match eq_ex (b8h ih) (b8h sh) with
+  [ true ⇒ match eq_ex (b8l ih) (b8l sh) with
+   [ true ⇒ match eq_ex (b8h il) (b8h sl) with
+    (* caso 0x...(X) a 0x...(Y) *)
+    [ true ⇒ setn_array16T (b8h ih) aux_2 data
+              (setn_array16T (b8l ih) aux_3 ilev2
+               (setn_array16T (b8h il) aux_4 ilev3
+                (* cambio a partire da livello 4 *)
+                (setmn_array16T (b8l il) (b8l sl) T ilev4 v))) (* ...X,...Y *)
+    (* caso 0x..(X1)(X2) a 0x..(Y1)(Y2) *)
+    | false ⇒ setn_array16T (b8h ih) aux_2 data
+               (setn_array16T (b8l ih) aux_3 ilev2
+                (* cambio a partire da livello 3 *)
+                (setn_array16T (b8h sl) aux_4 (* ..(Y1)0,..(Y1)(Y2) *)
+                 (setmn_array16T_succ_pred (b8h il) (b8h sl) aux_4 (* ..(X1+1).,..(Y1-1). *)
+                  (setn_array16T (b8h il) aux_4 ilev3
+                   (setmn_array16T (b8l il) xF T ilev4 v)) (* ..(X1)(X2),..(X1)F *)
+                  vlev4)
+                 (setmn_array16T x0 (b8l sl) T slev4 v))) ]
+   (* caso 0x.(X1)(X2)(X3) a 0x..(Y1)(Y2)(Y3) *)
+   | false ⇒ setn_array16T (b8h ih) aux_2 data
+              (* cambio a partire da livello 2 *)
+              (setn_array16T (b8l sh) aux_3
+               (setmn_array16T_succ_pred (b8l ih) (b8l sh) aux_3 (* .(X1+1)..,.(Y1-1).. *)
+                (setn_array16T (b8l ih) aux_3 ilev2
+                 (setmn_array16T_succ (b8h il) aux_4 (* .(X1)(X2+1).,.(X1)F. *)
+                  (setn_array16T (b8h il) aux_4 ilev3
+                   (setmn_array16T (b8l il) xF T ilev4 v)) (* .(X1)(X2)(X3),.(X1)(X2)F *)
+                  vlev4))
+                vlev3)
+               (setmn_array16T_pred (b8h sl) aux_4 (* .(Y1)0.,.(Y1)(Y2-1). *)
+                (setn_array16T (b8h sl) aux_4 slev3
+                 (setmn_array16T x0 (b8l sl) T slev4 v)) (* .(Y1)(Y2)0,.(Y1)(Y2)(Y3) *)
+                vlev4))
+   ]
+  (* caso 0x(X1)(X2)(X3)(X4) a 0x(Y1)(Y2)(Y3)(Y4) *)
+  | false ⇒ setn_array16T (b8h sh) aux_2
+             (setmn_array16T_succ_pred (b8h ih) (b8h sh) aux_2 (* (X+1)...,(Y-1)... *)
+              (setn_array16T (b8h ih) aux_2 data
+               (setmn_array16T_succ (b8l ih) aux_3 (* (X1)(X2+1)..,(X1)F.. *)
+                (setn_array16T (b8l ih) aux_3 ilev2
+                 (setmn_array16T_succ (b8h il) aux_4 (* (X1)(X2)(X3+1).,(X1)(X2)F. *)
+                  (setn_array16T (b8h il) aux_4 ilev3
+                   (setmn_array16T (b8l il) xF T ilev4 v)) (* (X1)(X2)(X3)(X4),(X1)(X2)(X3)F *)
+                  vlev4))
+                vlev3))
+               vlev2)
+             (setmn_array16T_pred (b8l sh) aux_3 (* (Y1)0..,(Y1)(Y2-1).. *)
+              (setn_array16T (b8l sh) aux_3 slev2
+               (setmn_array16T_pred (b8h sl) aux_4 (* (Y1)(Y2)0.,(Y1)(Y2)(Y3-1). *)
+                (setn_array16T (b8h sl) aux_4 slev3
+                 (setmn_array16T x0 (b8l sl) T slev4 v)) (* (Y1)(Y2)(Y3)0,(Y1)(Y2)(Y3)(Y4) *)
+                vlev4))
+              vlev3) ]]]
+ (* no i>s *)
+ | false ⇒ data
+ ].
+
+(* scrivi controllando il tipo di memoria *)
+definition mt_mem_update ≝
+λmem:Prod16T (Prod16T (Prod16T (Prod16T byte8))).
+λchk:Prod16T (Prod16T (Prod16T (Prod16T memory_type))).
+λaddr:word16.λv:byte8.
+ match mt_visit ? chk addr with
+  (* ROM? ok, ma il valore viene perso *)
+  [ MEM_READ_ONLY ⇒ Some ? mem
+  (* RAM? ok *)
+  | MEM_READ_WRITE ⇒ Some ? (mt_update ? mem addr v)
+  (* NON INSTALLATA? no *)
+  | MEM_OUT_OF_BOUND ⇒ None ? ].  
+
+(* leggi controllando il tipo di memoria *)
+definition mt_mem_read ≝
+λmem:Prod16T (Prod16T (Prod16T (Prod16T byte8))).
+λchk:Prod16T (Prod16T (Prod16T (Prod16T memory_type))).
+λaddr:word16.
+ match  mt_visit ? chk addr with
+  [ MEM_READ_ONLY ⇒ Some ? (mt_visit ? mem addr)
+  | MEM_READ_WRITE ⇒ Some ? (mt_visit ? mem addr)
+  | MEM_OUT_OF_BOUND ⇒ None ? ].
+
+(* ************************** *)
+(* CARICAMENTO PROGRAMMA/DATI *)
+(* ************************** *)
+
+(* carica a paratire da addr, scartando source (pescando da old_mem) se si supera 0xFFFF... *)
+let rec mt_load_from_source_at (old_mem:Prod16T (Prod16T (Prod16T (Prod16T byte8))))
+                               (source:list byte8) (addr:word16) on source ≝
+match source with
+ (* fine di source: carica da old_mem *)
+ [ nil ⇒ old_mem
+ | cons hd tl ⇒ match lt_w16 addr 〈〈xF,xF〉:〈xF,xF〉〉 with
+  (* non supera 0xFFFF, ricorsione *)
+  [ true ⇒ mt_load_from_source_at (mt_update ? old_mem addr hd) tl (plus_w16nc addr 〈〈x0,x0〉:〈x0,x1〉〉)
+  (* supera 0xFFFF, niente ricorsione *)
+  | false ⇒ mt_update ? old_mem addr hd
+  ]].
+
+(* ********************** *)
+(* TEOREMI/LEMMMI/ASSIOMI *)
+(* ********************** *)
+
+(* lettura da locazione diversa da quella in cui scrivo e' vecchio *)
+(* NB: sembra corretto, ma 2^32 casi!!! *)
+(*
+lemma ok_mt_update :
+ forall_word16 (λaddr1.
+ forall_word16 (λaddr2.
+  (eq_w16 addr1 addr2) ⊙
+  (eq_b8 (mt_visit ? (mt_update ? mt_zero_memory addr1 〈xF,xF〉) addr2)
+         (mk_byte8 x0 x0))))
+  = true.
+ reflexivity.
+qed.
+*)
+
+(* verifica scrittura di range *)
+(* NB: sembra corretto, ma 2^48 casi!!! *)
+(*
+lemma ok_mt_update_ranged :
+ forall_word16 (λaddr1.
+ forall_word16 (λaddr2.
+ forall_word16 (λaddr3.
+  (in_range addr3 addr1 addr2) ⊙
+  (eq_b8 (mt_visit ? (mt_update_ranged ? mt_zero_memory addr1 addr2 〈xF,xF〉) addr3)
+         (mk_byte8 x0 x0)))))
+  = true.
+ reflexivity.
+qed.
+*)
diff --git a/helm/software/matita/library/freescale/micro_tests.ma b/helm/software/matita/library/freescale/micro_tests.ma
new file mode 100644 (file)
index 0000000..a9c4797
--- /dev/null
@@ -0,0 +1,774 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/micro_tests/".
+
+(*include "/media/VIRTUOSO/freescale/multivm.ma".*)
+include "freescale/multivm.ma".
+
+(* ****************************************** *)
+(* MICRO TEST DI CORRETTEZZA DELLE ISTRUZIONI *)
+(* ****************************************** *)
+
+(* tabella 0x00 - 0xFF utile da caricare in RAM/ROM *)
+definition mTest_bytes : list byte8 ≝
+ [   〈x0,x0〉 ; 〈x0,x1〉 ; 〈x0,x2〉 ; 〈x0,x3〉 ; 〈x0,x4〉 ; 〈x0,x5〉 ; 〈x0,x6〉 ; 〈x0,x7〉
+ ; 〈x0,x8〉 ; 〈x0,x9〉 ; 〈x0,xA〉 ; 〈x0,xB〉 ; 〈x0,xC〉 ; 〈x0,xD〉 ; 〈x0,xE〉 ; 〈x0,xF〉 ]
+@[   〈x1,x0〉 ; 〈x1,x1〉 ; 〈x1,x2〉 ; 〈x1,x3〉 ; 〈x1,x4〉 ; 〈x1,x5〉 ; 〈x1,x6〉 ; 〈x1,x7〉
+ ; 〈x1,x8〉 ; 〈x1,x9〉 ; 〈x1,xA〉 ; 〈x1,xB〉 ; 〈x1,xC〉 ; 〈x1,xD〉 ; 〈x1,xE〉 ; 〈x1,xF〉 ]
+@[   〈x2,x0〉 ; 〈x2,x1〉 ; 〈x2,x2〉 ; 〈x2,x3〉 ; 〈x2,x4〉 ; 〈x2,x5〉 ; 〈x2,x6〉 ; 〈x2,x7〉
+ ; 〈x2,x8〉 ; 〈x2,x9〉 ; 〈x2,xA〉 ; 〈x2,xB〉 ; 〈x2,xC〉 ; 〈x2,xD〉 ; 〈x2,xE〉 ; 〈x2,xF〉 ]
+@[   〈x3,x0〉 ; 〈x3,x1〉 ; 〈x3,x2〉 ; 〈x3,x3〉 ; 〈x3,x4〉 ; 〈x3,x5〉 ; 〈x3,x6〉 ; 〈x3,x7〉
+ ; 〈x3,x8〉 ; 〈x3,x9〉 ; 〈x3,xA〉 ; 〈x3,xB〉 ; 〈x3,xC〉 ; 〈x3,xD〉 ; 〈x3,xE〉 ; 〈x3,xF〉 ]
+@[   〈x4,x0〉 ; 〈x4,x1〉 ; 〈x4,x2〉 ; 〈x4,x3〉 ; 〈x4,x4〉 ; 〈x4,x5〉 ; 〈x4,x6〉 ; 〈x4,x7〉
+ ; 〈x4,x8〉 ; 〈x4,x9〉 ; 〈x4,xA〉 ; 〈x4,xB〉 ; 〈x4,xC〉 ; 〈x4,xD〉 ; 〈x4,xE〉 ; 〈x4,xF〉 ]
+@[   〈x5,x0〉 ; 〈x5,x1〉 ; 〈x5,x2〉 ; 〈x5,x3〉 ; 〈x5,x4〉 ; 〈x5,x5〉 ; 〈x5,x6〉 ; 〈x5,x7〉
+ ; 〈x5,x8〉 ; 〈x5,x9〉 ; 〈x5,xA〉 ; 〈x5,xB〉 ; 〈x5,xC〉 ; 〈x5,xD〉 ; 〈x5,xE〉 ; 〈x5,xF〉 ]
+@[   〈x6,x0〉 ; 〈x6,x1〉 ; 〈x6,x2〉 ; 〈x6,x3〉 ; 〈x6,x4〉 ; 〈x6,x5〉 ; 〈x6,x6〉 ; 〈x6,x7〉
+ ; 〈x6,x8〉 ; 〈x6,x9〉 ; 〈x6,xA〉 ; 〈x6,xB〉 ; 〈x6,xC〉 ; 〈x6,xD〉 ; 〈x6,xE〉 ; 〈x6,xF〉 ]
+@[   〈x7,x0〉 ; 〈x7,x1〉 ; 〈x7,x2〉 ; 〈x7,x3〉 ; 〈x7,x4〉 ; 〈x7,x5〉 ; 〈x7,x6〉 ; 〈x7,x7〉
+ ; 〈x7,x8〉 ; 〈x7,x9〉 ; 〈x7,xA〉 ; 〈x7,xB〉 ; 〈x7,xC〉 ; 〈x7,xD〉 ; 〈x7,xE〉 ; 〈x7,xF〉 ]
+@[  〈x8,x0〉 ; 〈x8,x1〉 ; 〈x8,x2〉 ; 〈x8,x3〉 ; 〈x8,x4〉 ; 〈x8,x5〉 ; 〈x8,x6〉 ; 〈x8,x7〉
+ ; 〈x8,x8〉 ; 〈x8,x9〉 ; 〈x8,xA〉 ; 〈x8,xB〉 ; 〈x8,xC〉 ; 〈x8,xD〉 ; 〈x8,xE〉 ; 〈x8,xF〉 ]
+@[   〈x9,x0〉 ; 〈x9,x1〉 ; 〈x9,x2〉 ; 〈x9,x3〉 ; 〈x9,x4〉 ; 〈x9,x5〉 ; 〈x9,x6〉 ; 〈x9,x7〉
+ ; 〈x9,x8〉 ; 〈x9,x9〉 ; 〈x9,xA〉 ; 〈x9,xB〉 ; 〈x9,xC〉 ; 〈x9,xD〉 ; 〈x9,xE〉 ; 〈x9,xF〉 ]
+@[   〈xA,x0〉 ; 〈xA,x1〉 ; 〈xA,x2〉 ; 〈xA,x3〉 ; 〈xA,x4〉 ; 〈xA,x5〉 ; 〈xA,x6〉 ; 〈xA,x7〉
+ ; 〈xA,x8〉 ; 〈xA,x9〉 ; 〈xA,xA〉 ; 〈xA,xB〉 ; 〈xA,xC〉 ; 〈xA,xD〉 ; 〈xA,xE〉 ; 〈xA,xF〉 ]
+@[   〈xB,x0〉 ; 〈xB,x1〉 ; 〈xB,x2〉 ; 〈xB,x3〉 ; 〈xB,x4〉 ; 〈xB,x5〉 ; 〈xB,x6〉 ; 〈xB,x7〉
+ ; 〈xB,x8〉 ; 〈xB,x9〉 ; 〈xB,xA〉 ; 〈xB,xB〉 ; 〈xB,xC〉 ; 〈xB,xD〉 ; 〈xB,xE〉 ; 〈xB,xF〉 ]
+@[   〈xC,x0〉 ; 〈xC,x1〉 ; 〈xC,x2〉 ; 〈xC,x3〉 ; 〈xC,x4〉 ; 〈xC,x5〉 ; 〈xC,x6〉 ; 〈xC,x7〉
+ ; 〈xC,x8〉 ; 〈xC,x9〉 ; 〈xC,xA〉 ; 〈xC,xB〉 ; 〈xC,xC〉 ; 〈xC,xD〉 ; 〈xC,xE〉 ; 〈xC,xF〉 ]
+@[   〈xD,x0〉 ; 〈xD,x1〉 ; 〈xD,x2〉 ; 〈xD,x3〉 ; 〈xD,x4〉 ; 〈xD,x5〉 ; 〈xD,x6〉 ; 〈xD,x7〉
+ ; 〈xD,x8〉 ; 〈xD,x9〉 ; 〈xD,xA〉 ; 〈xD,xB〉 ; 〈xD,xC〉 ; 〈xD,xD〉 ; 〈xD,xE〉 ; 〈xD,xF〉 ]
+@[   〈xE,x0〉 ; 〈xE,x1〉 ; 〈xE,x2〉 ; 〈xE,x3〉 ; 〈xE,x4〉 ; 〈xE,x5〉 ; 〈xE,x6〉 ; 〈xE,x7〉
+ ; 〈xE,x8〉 ; 〈xE,x9〉 ; 〈xE,xA〉 ; 〈xE,xB〉 ; 〈xE,xC〉 ; 〈xE,xD〉 ; 〈xE,xE〉 ; 〈xE,xF〉 ]
+@[   〈xF,x0〉 ; 〈xF,x1〉 ; 〈xF,x2〉 ; 〈xF,x3〉 ; 〈xF,x4〉 ; 〈xF,x5〉 ; 〈xF,x6〉 ; 〈xF,x7〉
+ ; 〈xF,x8〉 ; 〈xF,x9〉 ; 〈xF,xA〉 ; 〈xF,xB〉 ; 〈xF,xC〉 ; 〈xF,xD〉 ; 〈xF,xE〉 ; 〈xF,xF〉
+ ].
+
+(* RIASSUNTO MCU/MODELLO/MEMORIA DA USARE
+
+   HC05 (MC68HC05J5A)   [0x0080-0x00FF] RAM [0x0300-0x0CFF] ROM
+   HC08 (MC68HC08AB16A) [0x0050-0x024F] RAM [0xBE00-0xFDFF] ROM
+   HCS08 (MC9S08AW60)   [0x0070-0x086F] RAM [0x1860-0xFFFF] FLASH
+   RS08 (MC9RS08KA2)    [0x0020-0x004F] RAM [0x3800-0x3FFF] FLASH *)
+
+(*
+   1) mTest_x_RAM : inizio della RAM
+       (start point per caricamento mTest_bytes in RAM) 
+   2) mTest_x_prog: inizio della ROM
+       (start point per caricamento programma in ROM)
+   3) mTest_x_data: ultimi 256b della ROM
+       (start point per caricamento mTest_bytes in ROM)
+*)
+definition mTest_HC05_RAM ≝ 〈〈x0,x0〉:〈x8,x0〉〉.
+definition mTest_HC05_prog ≝ 〈〈x0,x3〉:〈x0,x0〉〉.
+definition mTest_HC05_data ≝ 〈〈x0,xC〉:〈x0,x0〉〉.
+
+definition mTest_HC08_RAM ≝ 〈〈x0,x0〉:〈x5,x0〉〉.
+definition mTest_HC08_prog ≝ 〈〈xB,xE〉:〈x0,x0〉〉.
+definition mTest_HC08_data ≝ 〈〈xF,xD〉:〈x0,x0〉〉.
+
+definition mTest_HCS08_RAM ≝ 〈〈x0,x0〉:〈x7,x0〉〉.
+definition mTest_HCS08_prog ≝ 〈〈x1,x8〉:〈x6,x0〉〉.
+definition mTest_HCS08_data ≝ 〈〈xF,xF〉:〈x0,x0〉〉.
+
+definition mTest_RS08_RAM ≝ 〈〈x0,x0〉:〈x2,x0〉〉.
+definition mTest_RS08_prog ≝ 〈〈x3,x8〉:〈x0,x0〉〉.
+definition mTest_RS08_data ≝ 〈〈x3,xF〉:〈x0,x0〉〉.
+
+(* ********* *)
+(* HCS08 ADC *)
+(* ********* *)
+
+(* testa la logica di ADC e le modalita' IMM1,DIR1/2,IX0/1/2,SP1/2 *)
+definition mTest_HCS08_ADC_source : list byte8 ≝
+let m ≝ HCS08 in source_to_byte8 m (
+(* BEFORE: A=0x00 H:X=0xFF50 PC=0x1860 SP=0x0110 C=true *)
+(* [0x1860] 2clk *) (compile m ? ADC (maIMM1 〈xA,xA〉) I) @
+(* AFTER1: imm1=0xAA quindi 0x00+imm1+true=A:0xAB C:false *)
+(* [0x1862] 3clk *) (compile m ? ADC (maDIR1 〈xF,xF〉) I) @
+(* AFTER2: dir1=[0x00FF]=0x8F quindi 0xAB+dir1+false=A:0x3A C:true *)
+(* [0x1864] 4clk *) (compile m ? ADC (maDIR2 〈〈xF,xF〉:〈x1,x1〉〉) I) @
+(* AFTER3: dir2=[0xFF11]=0x11 quindi 0x3A+dir2+true=A:0x4C C:false *)
+(* [0x1867] 4clk *) (compile m ? ADC (maIX2 〈〈xF,xF〉:〈xF,x0〉〉) I) @
+(* AFTER4: ix2=[X+0xFFF0]=[0xFF40]=0x40 quindi 0x4C+ix2+false=A:0x8C C:false *)
+(* [0x186A] 3clk *) (compile m ? ADC (maIX1 〈x2,x4〉) I) @
+(* AFTER5: ix1=[X+0x0024]=[0xFF74]=0x74 quindi 0x8C+ix1+false=A:0x00 C:true *)
+(* [0x186C] 3clk *) (compile m ? ADC maIX0 I) @
+(* AFTER6: ix0=[X]=[0xFF50]=0x50 quindi 0x00+ix0+true=A:0x51 C:false *)
+(* [0x186D] 5clk *) (compile m ? ADC (maSP2 〈〈xF,xF〉:〈x6,x1〉〉) I) @
+(* AFTER7: sp2=[SP+0xFF61]=[0x0071]=0x01 quindi 0x51+sp2+false=A:0x52 C:false *)
+(* [0x1871] 4clk *) (compile m ? ADC (maSP1 〈x2,x4〉) I)
+(* AFTER8: sp1=[SP+0x0024]=[0x0134]=0xC4 quindi 0x52+sp1+false=A:0x16 C:true *)
+(* [0x1874] si puo' quindi enunciare che dopo 2+3+4+4+3+3+5+4=28 clk *)
+(*          A<-0x16 PC<-0x1874 *)
+ ).
+
+(* creazione del processore+caricamento+impostazione registri *)
+definition mTest_HCS08_ADC_status ≝
+λt:memory_impl.
+ set_c_flag HCS08 t (* C<-true *)
+  (setweak_sp_reg HCS08 t (* SP<-0x0110 *)
+   (setweak_indX_16_reg HCS08 t (* H:X<-0xFF50 *)
+    (set_pc_reg HCS08 t (* PC<-mTest_HCS08_prog *)
+     (start_of_mcu_version
+      MC9S08AW60 t
+      (load_from_source_at t (* carica mTest_bytes in ROM:mTest_HCS08_data *)
+       (load_from_source_at t (* carica mTest_bytes in RAM:mTest_HCS08_RAM *)
+        (load_from_source_at t (zero_memory t) (* carica source in ROM:mTest_HCS08_prog *)
+          mTest_HCS08_ADC_source mTest_HCS08_prog)
+         mTest_bytes mTest_HCS08_RAM)
+        mTest_bytes mTest_HCS08_data)
+      (build_memory_type_of_mcu_version MC9S08AW60 t)
+      (mk_byte8 x0 x0) (mk_byte8 x0 x0) (* non deterministici tutti a 0 *)
+      false false false false false false) (* non deterministici tutti a 0 *)
+     mTest_HCS08_prog)
+    (mk_word16 (mk_byte8 xF xF) (mk_byte8 x5 x0)))
+   (mk_word16 (mk_byte8 x0 x1) (mk_byte8 x1 x0)))
+  true.
+
+(* dimostrazione senza svolgimento degli stati, immediata *)
+lemma ok_mTest_HCS08_ADC_full :
+ ∀t:memory_impl.
+ execute HCS08 t (TickOK ? (mTest_HCS08_ADC_status t)) 28 =
+ (* NB: V,N,Z sono tornati false C e' tornato true *)
+ TickOK ? (set_pc_reg HCS08 t (* nuovo PC *)
+           (set_acc_8_low_reg HCS08 t (mTest_HCS08_ADC_status t) 〈x1,x6〉) (* nuovo A *)
+            (mk_word16 〈x1,x8〉 〈x7,x4〉)).
+ intro;
+ elim t;
+ reflexivity.
+qed.
+
+(* ********* *)
+(* HCS08 MOV *)
+(* ********* *)
+
+(* testa la logica di MOV e le modalita' xxx_to_xxx *)
+definition mTest_HCS08_MOV_source : list byte8 ≝
+let m ≝ HCS08 in source_to_byte8 m (
+(* BEFORE: H:X=0x0071 PC=0x1860 *)
+
+(* [0x1860] 4clk *) (compile m ? MOV (maIMM1_to_DIR1 〈x3,xA〉 〈x7,x0〉) I) @
+(* 0x3A in [0x0070] *)
+(* [0x1863] 5clk *) (compile m ? MOV (maDIR1_to_DIR1 〈x7,x0〉 〈x7,x1〉) I) @
+(* [0x0070] in [0x0071] *)
+(* [0x1866] 5clk *) (compile m ? MOV (maIX0p_to_DIR1 〈x7,x2〉) I) @
+(* [X++=0x0071] in [0x0072] *)
+(* [0x1868] 5clk *) (compile m ? MOV (maDIR1_to_IX0p 〈x7,x2〉) I)
+(* [0x0072] in [X++=0x0072] *)
+(* [0x186A] si puo' quindi enunciare che dopo 4+5+5+5=19 clk *)
+(*          PC<-0x186A X<-0x0073 *)
+ ).
+
+(* creazione del processore+caricamento+impostazione registri *)
+definition mTest_HCS08_MOV_status ≝
+λt:memory_impl.
+λb1,b2,b3:byte8.
+ setweak_indX_16_reg HCS08 t (* H:X<-0x0071 *)
+  (set_pc_reg HCS08 t (* PC<-mTest_HCS08_prog *)
+   (start_of_mcu_version
+    MC9S08AW60 t
+    (load_from_source_at t (* carica b1-3 in RAM:mTest_HCS08_RAM *)
+     (load_from_source_at t (zero_memory t) (* carica source in ROM:mTest_HCS08_prog *)
+       mTest_HCS08_MOV_source mTest_HCS08_prog)
+     [ b1 ; b2 ; b3 ] mTest_HCS08_RAM)
+    (build_memory_type_of_mcu_version MC9S08AW60 t)
+    (mk_byte8 x0 x0) (mk_byte8 x0 x0) (* non deterministici tutti a 0 *)
+    false false false false false false) (* non deterministici tutti a 0 *)
+   mTest_HCS08_prog)
+  (mk_word16 (mk_byte8 x0 x0) (mk_byte8 x7 x1)).
+
+(* dimostrazione senza svolgimento degli stati, immediata *)
+(* NB: la memoria e' cambiata e bisogna applicare eq_status_axiom *)
+lemma ok_mTest_HCS08_MOV_full :
+ ∀t:memory_impl.
+ match execute HCS08 t (TickOK ? (mTest_HCS08_MOV_status t 〈x0,x0〉 〈x0,x0〉 〈x0,x0〉)) 19 with 
+  [ TickERR s _ ⇒ s | TickSUSP s _ ⇒ s | TickOK s ⇒ s ] =
+ setweak_indX_16_reg HCS08 t (* H:X *) 
+  (set_pc_reg HCS08 t (mTest_HCS08_MOV_status t 〈x3,xA〉 〈x3,xA〉 〈x3,xA〉) (* PC *)
+   (mk_word16 〈x1,x8〉 〈x6,xA〉))
+    (mk_word16 〈x0,x0〉 〈x7,x3〉).
+ intro;
+ elim t;
+ [ apply (eq_status_axiom mTest_HCS08_RAM 2 HCS08 MEM_FUNC); reflexivity; ] 
+ reflexivity.
+qed.
+
+(* ************* *)
+(* HCS08 ROL/ROR *)
+(* ************* *)
+
+(* testa la logica di ROL/ROR e le modalita' IMM2,INHx *)
+definition mTest_HCS08_ROL_ROR_source : list byte8 ≝
+let m ≝ HCS08 in source_to_byte8 m (
+(* BEFORE: A=0x00 H:X=0x0000 PC=0x1860 Z=true *)
+(* [0x1860] 3clk *) (compile m ? LDHX (maIMM2 〈〈x1,x2〉:〈x3,x4〉〉) I) @
+(* [0x1863] 2clk *) (compile m ? LDA (maIMM1 〈x5,x6〉) I) @
+(* [0x1865] 1clk *) (compile m ? ROL maINHA I) @
+(* [0x1866] 1clk *) (compile m ? ROL maINHX I) @
+(* [0x1867] 1clk *) (compile m ? ROR maINHA I) @
+(* [0x1868] 1clk *) (compile m ? ROR maINHX I) @
+(* [0x1869] 1clk *) (compile m ? CLR maINHA I) @
+(* [0x186A] 1clk *) (compile m ? CLR maINHX I) @
+(* [0x186B] 1clk *) (compile m ? CLR maINHH I)
+(* [0x186C] si puo' quindi enunciare che dopo 3+2+1+1+1+1+1+1+1=12 clk *)
+(*          PC<-0x186C *)
+ ).
+
+(* creazione del processore+caricamento+impostazione registri *)
+definition mTest_HCS08_ROL_ROR_status ≝
+λt:memory_impl.
+ set_z_flag HCS08 t (* Z<-true *)
+  (setweak_indX_16_reg HCS08 t (* H:X<-0x0000 *)
+   (set_pc_reg HCS08 t (* PC<-mTest_HCS08_prog *)
+    (start_of_mcu_version
+     MC9S08AW60 t
+     (load_from_source_at t (zero_memory t) (* carica source in ROM:mTest_HCS08_prog *)
+      mTest_HCS08_ROL_ROR_source mTest_HCS08_prog)
+     (build_memory_type_of_mcu_version MC9S08AW60 t)
+     (mk_byte8 x0 x0) (mk_byte8 x0 x0) (* non deterministici tutti a 0 *)
+     false false false false false false) (* non deterministici tutti a 0 *)
+    mTest_HCS08_prog)
+   (mk_word16 (mk_byte8 x0 x0) (mk_byte8 x0 x0)))
+  true.
+
+(* dimostrazione senza svolgimento degli stati, immediata *)
+lemma ok_mTest_HCS08_ROL_ROR_full :
+ ∀t:memory_impl.
+ execute HCS08 t (TickOK ? (mTest_HCS08_ROL_ROR_status t)) 12 =
+ TickOK ? (set_pc_reg HCS08 t (* nuovo PC *)
+           (mTest_HCS08_ROL_ROR_status t) (mk_word16 〈x1,x8〉 〈x6,xC〉)).
+ intro;
+ elim t;
+ reflexivity.
+qed.
+
+(* **************** *)
+(* HCS08 CBEQx/DBNZ *)
+(* **************** *)
+
+(* testa la logica di CBEQx/DBNZ e le modalita' xxx_and_IMM1 *)
+definition mTest_HCS08_CBEQ_DBNZ_source : list byte8 ≝
+let m ≝ HCS08 in source_to_byte8 m (
+(* BEFORE: H:X=0x006F SP=0x006F PC=0x1860 *)
+(* [0x1860] 5clk *) (compile m ? CBEQA (maDIR1_and_IMM1 〈x7,x1〉 〈x0,x1〉) I) @
+(* [0x1863] 1clk *) (compile m ? NOP maINH I) @ (* eseguito: A≠[0x0071]=0x01 *)
+(* [0x1864] 4clk *) (compile m ? CBEQA (maIMM1_and_IMM1 〈x0,x0〉 〈x0,x1〉) I) @
+(* [0x1867] 1clk *) (compile m ? NOP maINH I) @ (* non eseguito: A=0x00 *)
+(* [0x1868] 4clk *) (compile m ? CBEQX (maIMM1_and_IMM1 〈x6,xF〉 〈x0,x1〉) I) @
+(* [0x186B] 1clk *) (compile m ? NOP maINH I) @ (* non eseguito: X=0x6F *)
+(* [0x186C] 5clk *) (compile m ? CBEQA (maIX1p_and_IMM1 〈x0,x1〉 〈x0,x1〉) I) @ (* H:X++ *)
+(* [0x186F] 1clk *) (compile m ? NOP maINH I) @ (* non eseguito: A=[X+0x01]=[0x0070]=0x00 *)
+(* [0x1870] 5clk *) (compile m ? CBEQA (maIX0p_and_IMM1 〈x0,x1〉) I) @ (* H:X++ *)
+(* [0x1872] 1clk *) (compile m ? NOP maINH I) @ (* non eseguito: A=[X]=[0x0070]=0x00 *)
+(* [0x1873] 6clk *) (compile m ? CBEQA (maSP1_and_IMM1 〈x0,x2〉 〈x0,x1〉) I) @
+(* [0x1877] 1clk *) (compile m ? NOP maINH I) @ (* eseguito: A≠[SP+0x02]=[0x0071]=0x01 *)
+
+(* [0x1878] 7clk *) (compile m ? DBNZ (maDIR1_and_IMM1 〈x7,x2〉 〈x0,x1〉) I) @
+(* [0x187B] 1clk *) (compile m ? NOP maINH I) @ (* non eseguito: --[0x0072]=0x01≠0 *)
+(* [0x187C] 4clk *) (compile m ? DBNZ (maINHA_and_IMM1 〈x0,x1〉) I) @
+(* [0x187E] 1clk *) (compile m ? NOP maINH I) @ (* non eseguito: --A=0xFF≠0 *)
+(* [0x187F] 4clk *) (compile m ? DBNZ (maINHX_and_IMM1 〈x0,x1〉) I) @
+(* [0x1881] 1clk *) (compile m ? NOP maINH I) @ (* non eseguito: --X=0x70≠0 *)
+(* [0x1882] 7clk *) (compile m ? DBNZ (maIX1_and_IMM1 〈x0,x2〉 〈x0,x1〉) I) @
+(* [0x1885] 1clk *) (compile m ? NOP maINH I) @ (* eseguito: --[X+0x02]=[0x0072]=0x00=0 *)
+(* [0x1886] 6clk *) (compile m ? DBNZ (maIX0_and_IMM1 〈x0,x1〉) I) @
+(* [0x1888] 1clk *) (compile m ? NOP maINH I) @ (* non eseguito: --[X]=[0x0070]=0xFF≠0 *)
+(* [0x1889] 8clk *) (compile m ? DBNZ (maSP1_and_IMM1 〈x0,x1〉 〈x0,x1〉) I) @ 
+(* [0x188D] 1clk *) (compile m ? NOP maINH I) (* non eseguito: --[SP+0x01]=[0x0070]=0xFE≠0 *)
+(* [0x188E] si puo' quindi enunciare che dopo 5+1+4+4+5+5+6+1 (31)
+                                              7+4+4+7+1+6+8   (37) =68 clk *)
+(*          A<-0xFF PC<-0x188E H:X<-0070 *)
+ ).
+
+(* creazione del processore+caricamento+impostazione registri *)
+definition mTest_HCS08_CBEQ_DBNZ_status ≝
+λt:memory_impl.
+λb1,b2,b3:byte8.
+ setweak_sp_reg HCS08 t (* SP<-0x006F *)
+  (setweak_indX_16_reg HCS08 t (* H:X<-0x006F *)
+   (set_pc_reg HCS08 t (* PC<-mTest_HCS08_prog *)
+    (start_of_mcu_version
+     MC9S08AW60 t
+     (load_from_source_at t (* carica b1-3 in RAM:mTest_HCS08_RAM *)
+      (load_from_source_at t (* carica mTest_bytes in RAM:mTest_HCS08_RAM *)
+       (load_from_source_at t (zero_memory t) (* carica source in ROM:mTest_HCS08_prog *)
+         mTest_HCS08_CBEQ_DBNZ_source mTest_HCS08_prog)
+        mTest_bytes mTest_HCS08_RAM)
+       [ b1 ; b2 ; b3 ] mTest_HCS08_RAM)
+     (build_memory_type_of_mcu_version MC9S08AW60 t)
+     (mk_byte8 x0 x0) (mk_byte8 x0 x0) (* non deterministici tutti a 0 *)
+     false false false false false false) (* non deterministici tutti a 0 *)
+    mTest_HCS08_prog)
+   (mk_word16 (mk_byte8 x0 x0) (mk_byte8 x6 xF)))
+  (mk_word16 (mk_byte8 x0 x0) (mk_byte8 x6 xF)).
+
+(* dimostrazione senza svolgimento degli stati, immediata *)
+(* NB: la memoria e' cambiata e bisogna applicare eq_status_axiom *)
+lemma ok_mTest_HCS08_CBEQ_DBNZ_full :
+ ∀t:memory_impl.
+ match execute HCS08 t (TickOK ? (mTest_HCS08_CBEQ_DBNZ_status t 〈x0,x0〉 〈x0,x1〉 〈x0,x2〉)) 68 with
+  [ TickERR s _ ⇒ s | TickSUSP s _ ⇒ s | TickOK s ⇒ s ] =
+ set_acc_8_low_reg HCS08 t (* nuovo A *)
+  (set_pc_reg HCS08 t (* nuovo PC *)
+   (setweak_indX_16_reg HCS08 t (mTest_HCS08_CBEQ_DBNZ_status t 〈xF,xE〉 〈x0,x1〉 〈x0,x0〉) (* nuovo H:X *)
+    (mk_word16 〈x0,x0〉 〈x7,x0〉))
+   (mk_word16 〈x1,x8〉 〈x8,xE〉))
+  (mk_byte8 xF xF).
+ intro;
+ elim t;
+ [ apply (eq_status_axiom mTest_HCS08_RAM 2 HCS08 MEM_FUNC); reflexivity; ]
+ reflexivity.
+qed.
+
+(* ***************** *)
+(* HCS08 BSETn/BCLRn *)
+(* ***************** *)
+
+(* testa la logica di BSETn/BCLRn e le modalita' DIRn *)
+definition mTest_HCS08_BSETn_BCLRn_source : list byte8 ≝
+let m ≝ HCS08 in source_to_byte8 m (
+(* BEFORE: PC=0x1860 *)
+(* [0x1860] 5clk *) (compile m ? BSETn (maDIRn o0 〈x7,x0〉) I) @ (* [0x0070]=0x01 *)
+(* [0x1862] 5clk *) (compile m ? BSETn (maDIRn o1 〈x7,x0〉) I) @ (* [0x0070]=0x03 *)
+(* [0x1864] 5clk *) (compile m ? BSETn (maDIRn o2 〈x7,x0〉) I) @ (* [0x0070]=0x07 *)
+(* [0x1866] 5clk *) (compile m ? BSETn (maDIRn o3 〈x7,x0〉) I) @ (* [0x0070]=0x0F *)
+(* [0x1868] 5clk *) (compile m ? BSETn (maDIRn o4 〈x7,x0〉) I) @ (* [0x0070]=0x1F *)
+(* [0x186A] 5clk *) (compile m ? BSETn (maDIRn o5 〈x7,x0〉) I) @ (* [0x0070]=0x3F *)
+(* [0x186C] 5clk *) (compile m ? BSETn (maDIRn o6 〈x7,x0〉) I) @ (* [0x0070]=0x7F *)
+(* [0x186E] 5clk *) (compile m ? BSETn (maDIRn o7 〈x7,x0〉) I) @ (* [0x0070]=0xFF *)
+(* [0x1870] 5clk *) (compile m ? BCLRn (maDIRn o0 〈x7,x0〉) I) @ (* [0x0070]=0xFE *)
+(* [0x1872] 5clk *) (compile m ? BCLRn (maDIRn o1 〈x7,x0〉) I) @ (* [0x0070]=0xFC *)
+(* [0x1874] 5clk *) (compile m ? BCLRn (maDIRn o2 〈x7,x0〉) I) @ (* [0x0070]=0xF8 *)
+(* [0x1876] 5clk *) (compile m ? BCLRn (maDIRn o3 〈x7,x0〉) I) @ (* [0x0070]=0xF0 *)
+(* [0x1878] 5clk *) (compile m ? BCLRn (maDIRn o4 〈x7,x0〉) I) @ (* [0x0070]=0xE0 *)
+(* [0x187A] 5clk *) (compile m ? BCLRn (maDIRn o5 〈x7,x0〉) I) @ (* [0x0070]=0xC0 *)
+(* [0x187C] 5clk *) (compile m ? BCLRn (maDIRn o6 〈x7,x0〉) I) @ (* [0x0070]=0x80 *)
+(* [0x187E] 5clk *) (compile m ? BCLRn (maDIRn o7 〈x7,x0〉) I)   (* [0x0070]=0x00 *)
+(* [0x1880] si puo' quindi enunciare che dopo 5+5+5+5+5+5+5+5
+                                              5+5+5+5+5+5+5+5 =80 clk *)
+(*          PC<-0x1880 *)
+ ).
+
+(* creazione del processore+caricamento+impostazione registri *)
+definition mTest_HCS08_BSETn_BCLRn_status ≝
+λt:memory_impl.
+λb1:byte8.
+ set_pc_reg HCS08 t (* PC<-mTest_HCS08_prog *)
+  (start_of_mcu_version
+   MC9S08AW60 t
+   (load_from_source_at t (* carica b1 in RAM:mTest_HCS08_RAM *)
+    (load_from_source_at t (zero_memory t) (* carica source in ROM:mTest_HCS08_prog *)
+      mTest_HCS08_BSETn_BCLRn_source mTest_HCS08_prog)
+     [ b1 ] mTest_HCS08_RAM)
+   (build_memory_type_of_mcu_version MC9S08AW60 t)
+   (mk_byte8 x0 x0) (mk_byte8 x0 x0) (* non deterministici tutti a 0 *)
+   false false false false false false) (* non deterministici tutti a 0 *)
+  mTest_HCS08_prog.
+
+(* dimostrazione senza svolgimento degli stati, immediata *)
+(* NB: la memoria e' cambiata e bisogna applicare eq_status_axiom *)
+lemma ok_mTest_HCS08_BSETn_BCLRn_full :
+ ∀t:memory_impl.
+ match execute HCS08 t (TickOK ? (mTest_HCS08_BSETn_BCLRn_status t 〈x0,x0〉)) 80 with
+  [ TickERR s _ ⇒ s | TickSUSP s _ ⇒ s | TickOK s ⇒ s ] =
+ set_pc_reg HCS08 t (mTest_HCS08_BSETn_BCLRn_status t 〈x0,x0〉) 〈〈x1,x8〉:〈x8,x0〉〉. (* nuovo PC *)
+ intro;
+ elim t;
+ [ apply (eq_status_axiom mTest_HCS08_RAM 0 HCS08 MEM_FUNC); reflexivity ]
+ reflexivity.
+qed.
+
+(* ******************* *)
+(* HCS08 BRSETn/BRCLRn *)
+(* ******************* *)
+
+(* testa la logica di BRSETn/BRCLRn e le modalita' DIRn_and_IMM1 *)
+definition mTest_HCS08_BRSETn_BRCLRn_source : list byte8 ≝
+let m ≝ HCS08 in source_to_byte8 m (
+(* BEFORE: va a testare [0x00C5]=0x55 PC=0x1860 *)
+(* [0x1860] 5clk *) (compile m ? BRSETn (maDIRn_and_IMM1 o0 〈xC,x5〉 〈x0,x1〉) I) @
+(* [0x1863] 1clk *) (compile m ? NOP maINH I) @
+(* [0x1864] 5clk *) (compile m ? BRSETn (maDIRn_and_IMM1 o1 〈xC,x5〉 〈x0,x1〉) I) @
+(* [0x1867] 1clk *) (compile m ? NOP maINH I) @
+(* [0x1868] 5clk *) (compile m ? BRSETn (maDIRn_and_IMM1 o2 〈xC,x5〉 〈x0,x1〉) I) @
+(* [0x186B] 1clk *) (compile m ? NOP maINH I) @
+(* [0x186C] 5clk *) (compile m ? BRSETn (maDIRn_and_IMM1 o3 〈xC,x5〉 〈x0,x1〉) I) @
+(* [0x186F] 1clk *) (compile m ? NOP maINH I) @
+(* [0x1870] 5clk *) (compile m ? BRSETn (maDIRn_and_IMM1 o4 〈xC,x5〉 〈x0,x1〉) I) @
+(* [0x1873] 1clk *) (compile m ? NOP maINH I) @
+(* [0x1874] 5clk *) (compile m ? BRSETn (maDIRn_and_IMM1 o5 〈xC,x5〉 〈x0,x1〉) I) @
+(* [0x1877] 1clk *) (compile m ? NOP maINH I) @
+(* [0x1878] 5clk *) (compile m ? BRSETn (maDIRn_and_IMM1 o6 〈xC,x5〉 〈x0,x1〉) I) @
+(* [0x187B] 1clk *) (compile m ? NOP maINH I) @
+(* [0x187C] 5clk *) (compile m ? BRSETn (maDIRn_and_IMM1 o7 〈xC,x5〉 〈x0,x1〉) I) @
+(* [0x187F] 1clk *) (compile m ? NOP maINH I) @
+
+(* [0x1880] 5clk *) (compile m ? BRCLRn (maDIRn_and_IMM1 o0 〈xC,x5〉 〈x0,x1〉) I) @
+(* [0x1883] 1clk *) (compile m ? NOP maINH I) @
+(* [0x1884] 5clk *) (compile m ? BRCLRn (maDIRn_and_IMM1 o1 〈xC,x5〉 〈x0,x1〉) I) @
+(* [0x1887] 1clk *) (compile m ? NOP maINH I) @
+(* [0x1888] 5clk *) (compile m ? BRCLRn (maDIRn_and_IMM1 o2 〈xC,x5〉 〈x0,x1〉) I) @
+(* [0x188B] 1clk *) (compile m ? NOP maINH I) @
+(* [0x188C] 5clk *) (compile m ? BRCLRn (maDIRn_and_IMM1 o3 〈xC,x5〉 〈x0,x1〉) I) @
+(* [0x188F] 1clk *) (compile m ? NOP maINH I) @
+(* [0x1890] 5clk *) (compile m ? BRCLRn (maDIRn_and_IMM1 o4 〈xC,x5〉 〈x0,x1〉) I) @
+(* [0x1893] 1clk *) (compile m ? NOP maINH I) @
+(* [0x1894] 5clk *) (compile m ? BRCLRn (maDIRn_and_IMM1 o5 〈xC,x5〉 〈x0,x1〉) I) @
+(* [0x1897] 1clk *) (compile m ? NOP maINH I) @
+(* [0x1898] 5clk *) (compile m ? BRCLRn (maDIRn_and_IMM1 o6 〈xC,x5〉 〈x0,x1〉) I) @
+(* [0x189B] 1clk *) (compile m ? NOP maINH I) @
+(* [0x189C] 5clk *) (compile m ? BRCLRn (maDIRn_and_IMM1 o7 〈xC,x5〉 〈x0,x1〉) I) @
+(* [0x189F] 1clk *) (compile m ? NOP maINH I)
+
+(* [0x18A0] si puo' quindi enunciare che dopo 80+8=88 clk
+            (vengono eseguiti 16*5 test, meta' BRSETn/BRCLRn saltano *) 
+(*          PC<-0x18A0 *)
+ ).
+
+(* creazione del processore+caricamento+impostazione registri *)
+definition mTest_HCS08_BRSETn_BRCLRn_status ≝
+λt:memory_impl.
+ set_pc_reg HCS08 t (* PC<-mTest_HCS08_prog *)
+  (start_of_mcu_version
+   MC9S08AW60 t
+   (load_from_source_at t
+    (load_from_source_at t (zero_memory t) (* carica mTest_bytes in RAM:mTest_HCS08_RAM *)
+      mTest_HCS08_BRSETn_BRCLRn_source mTest_HCS08_prog) (* carica source in ROM:mTest_HCS08_prog *)
+     mTest_bytes mTest_HCS08_RAM)
+   (build_memory_type_of_mcu_version MC9S08AW60 t)
+   (mk_byte8 x0 x0) (mk_byte8 x0 x0) (* non deterministici tutti a 0 *)
+   false false false false false false) (* non deterministici tutti a 0 *)
+  mTest_HCS08_prog.
+
+(* dimostrazione senza svolgimento degli stati, immediata *)
+lemma ok_mTest_HCS08_BRSETn_BRCLRn_full :
+ ∀t:memory_impl.
+ execute HCS08 t (TickOK ? (mTest_HCS08_BRSETn_BRCLRn_status t)) 88 =
+ TickOK ? (set_pc_reg HCS08 t (mTest_HCS08_BRSETn_BRCLRn_status t) (* nuovo PC *)
+           (mk_word16 〈x1,x8〉 〈xA,x0〉)).
+ intro;
+ elim t;
+ reflexivity.
+qed.
+
+(* *************** *)
+(* RS08 X,D[X],TNY *)
+(* *************** *)
+
+(* testa la logica RS08 X,D[X] le modalita' TNY *)
+(* NB: il meccanismo utilizzato e' quello complesso dell'RS08
+       fare riferimento alle spiegazioni in STATUS/LOAD_WRITE *)
+definition mTest_RS08_TNY_source: list byte8 ≝
+let m ≝ RS08 in source_to_byte8 m (
+(* X=20 PS=0 *)
+(* [0x3800] 3clk *) (compile m ? ADD (maTNY xD) I) @ (* ... +[0x000D]=0x0C *)
+(* [0x3801] 3clk *) (compile m ? ADD (maTNY xE) I) @ (* ... +D[X]=[0x0020]=0x1F *)
+(* [0x3802] 3clk *) (compile m ? ADD (maTNY xF) I) @ (* ... +X=0x20 *)
+(* [0x3803] 3clk *) (compile m ? ADD (maDIR1 〈xC,xF〉) I) @ (* ... +X=0x20 *)
+(* [0x3805] 3clk *) (compile m ? ADD (maDIR1 〈xC,xE〉) I) (* ... +[0x000E]=0x0D *)
+(* [0x3807] si puo' quindi enunciare che dopo 15 clk
+            A<-0x78 PC<-0x3807 *)
+ ).
+
+(* creazione del processore+caricamento+impostazione registri *)
+definition mTest_RS08_TNY_status ≝
+λt:memory_impl.
+ setweak_x_map RS08 t (* X<-0x20 *)
+ (setweak_ps_map RS08 t (* PS<-0x00 *)
+ (set_pc_reg RS08 t (* PC<-mTest_RS08_prog *)
+  (start_of_mcu_version 
+   MC9RS08KA2 t
+   (load_from_source_at t (* carica mTest_bytes in RAM:mTest_RS08_RAM *)
+    (load_from_source_at t (zero_memory t) (* carica source in ROM:mTest_RS08_prog *)
+     mTest_RS08_TNY_source mTest_RS08_prog)
+    mTest_bytes 〈〈x0,x0〉:〈x0,x1〉〉)
+   (build_memory_type_of_mcu_version MC9RS08KA2 t)
+   (mk_byte8 x0 x0) (mk_byte8 x0 x0) (* non deterministici tutti a 0 *)
+   false false false false false false (* non deterministici tutti a 0 *)
+   ) mTest_RS08_prog)
+  (mk_byte8 x0 x0))
+  (mk_byte8 x2 x0).
+
+(* dimostrazione senza svolgimento degli stati, immediata *)
+lemma ok_mTest_RS08_TNY_full :
+ ∀t:memory_impl.
+ execute RS08 t (TickOK ? (mTest_RS08_TNY_status t)) 15 =
+ TickOK ? (set_acc_8_low_reg RS08 t (* nuovo A *)
+           (set_pc_reg RS08 t (mTest_RS08_TNY_status t) (* nuovo PC *)
+            (mk_word16 〈x3,x8〉 〈x0,x7〉))
+             (mk_byte8 x7 x8)).
+ intro;
+ elim t;
+ reflexivity.
+qed.
+
+(* *********** *)
+(* RS08 PS,SRT *)
+(* *********** *)
+
+(* testa la logica RS08 PS le modalita' SRT *)
+(* NB: il meccanismo utilizzato e' quello complesso dell'RS08
+       fare riferimento alle spiegazioni in STATUS/LOAD_WRITE *)
+definition mTest_RS08_SRT_source: list byte8 ≝
+let m ≝ RS08 in source_to_byte8 m (
+(* X=0x1F PS=0xFE Z=1 *)
+(* [0x3800] 3clk *) (compile m ? LDA (maSRT t1F) I) @ (* A<-PS *)
+(* [0x3801] 2clk *) (compile m ? SUB (maIMM1 〈xF,xE〉) I) @ (* risulta 0 *)
+(* [0x3803] 3clk *) (compile m ? BEQ (maIMM1 〈x0,x1〉) I) @ (* salta *)
+(* [0x3805] 1clk *) (compile m ? NOP maINH I) @
+
+(* [0x3806] 3clk *) (compile m ? LDA (maSRT t0E) I) @ (* A<-PS *)
+(* [0x3807] 2clk *) (compile m ? SUB (maIMM1 〈xF,xE〉) I) @ (* risulta 0 *)
+(* [0x3809] 3clk *) (compile m ? BEQ (maIMM1 〈x0,x1〉) I) @ (* salta *)
+(* [0x380B] 1clk *) (compile m ? NOP maINH I) @
+
+(* [0x380C] 3clk *) (compile m ? LDA (maDIR1 〈xC,x3〉) I) @ (* A<-[0x00C3]=[0x3F83]=0x83 *)
+(* [0x380E] 2clk *) (compile m ? SUB (maIMM1 〈x8,x3〉) I) @ (* risulta 0 *)
+(* [0x3810] 3clk *) (compile m ? BEQ (maIMM1 〈x0,x1〉) I) @ (* salta *)
+(* [0x3812] 1clk *) (compile m ? NOP maINH I)
+(* [0x3813] si puo' quindi enunciare che dopo 24 clk
+            PC<-0x3813 *)
+ ).
+
+(* creazione del processore+caricamento+impostazione registri *)
+definition mTest_RS08_SRT_status ≝
+λt:memory_impl.
+ setweak_x_map RS08 t (* X<-0x1F *)
+  (setweak_ps_map RS08 t (* PS<-0xFE *)
+   (set_z_flag RS08 t (* Z<-true *)
+    (set_pc_reg RS08 t (* PC<-mTest_RS08_prog *)
+     (start_of_mcu_version 
+      MC9RS08KA2 t
+       (load_from_source_at t (* carica mTest_bytes in ROM:mTest_RS08_data *)
+        (load_from_source_at t (zero_memory t) (* carica source in ROM:mTest_RS08_prog *)
+         mTest_RS08_SRT_source mTest_RS08_prog)
+        mTest_bytes mTest_RS08_data)
+       (build_memory_type_of_mcu_version MC9RS08KA2 t)
+      (mk_byte8 x0 x0) (mk_byte8 x0 x0) (* non deterministici tutti a 0 *)
+      false false false false false false (* non deterministici tutti a 0 *)
+      ) mTest_RS08_prog)
+     true)
+    (mk_byte8 xF xE))
+   (mk_byte8 x1 xF).
+
+(* dimostrazione senza svolgimento degli stati, immediata *)
+lemma ok_mTest_RS08_SRT_full :
+ ∀t:memory_impl.
+ execute RS08 t (TickOK ? (mTest_RS08_SRT_status t)) 24 =
+ TickOK ? (set_pc_reg RS08 t (mTest_RS08_SRT_status t) (* nuovo PC *)
+           (mk_word16 〈x3,x8〉 〈x1,x3〉)).
+ intro;
+ elim t;
+ reflexivity.
+qed.
+
+(* ********************* *)
+(* TEOREMA MULT PER RS08 *)
+(* ********************* *)
+
+(* 
+   ZH:ZL=X*Y con [0x0020-0x004F] X ≝ [0x0020] Y ≝ [0x0021] ZH ≝ [0x0022] ZL ≝ [0x0023]
+*)
+definition mTest_RS08_mult_source: list byte8 ≝
+let m ≝ RS08 in source_to_byte8 m (
+(* [0x3800] ZH <- 0      3clk *) (compile m ? CLR (maDIR1 〈x2,x2〉) I) @
+(* [0x3802] ZL <- 0      3clk *) (compile m ? CLR (maDIR1 〈x2,x3〉) I) @
+(* [0x3804] (l1) A <- Y  3clk *) (compile m ? LDA (maDIR1 〈x2,x1〉) I) @
+(* [0x3806] A=0 goto l2  3clk *) (compile m ? BEQ (maIMM1 〈x0,xE〉) I) @
+(* [0x3808] A <- ZL      3clk *) (compile m ? LDA (maDIR1 〈x2,x3〉) I) @
+(* [0x380A] Y --         5clk *) (compile m ? DEC (maDIR1 〈x2,x1〉) I) @
+(* [0x380C] A += X       3clk *) (compile m ? ADD (maDIR1 〈x2,x0〉) I) @
+(* [0x380E] C=0 goto l3  3clk *) (compile m ? BCC (maIMM1 〈x0,x2〉) I) @
+(* [0x3810] ZH ++        5clk *) (compile m ? INC (maDIR1 〈x2,x2〉) I) @
+(* [0x3812] (l3) ZL <- A 3clk *) (compile m ? STA (maDIR1 〈x2,x3〉) I) @
+(* [0x3814] goto l1      3clk *) (compile m ? BRA (maIMM1 〈xE,xE〉) I)
+(* [0x3816] (l2) si puo' quindi enunciare che
+                 - il caso base X * 0 richiede 12 cicli
+                 - bisogna aggiungere Y * 26 cicli, Y>0\r
+                 - bisogna aggiungere ZH * 5 cicli, X * Y > 0xFF *)
+).
+
+(* creazione del processore+caricamento+impostazione registri *)
+definition mTest_RS08_mult_status ≝
+λt:memory_impl.
+λb1,b2,b3,b4:byte8.
+ set_z_flag RS08 t (* Z<-true *)
+ (set_pc_reg RS08 t (* PC<-mTest_RS08_prog *)
+  (start_of_mcu_version 
+   MC9RS08KA2 t
+   (load_from_source_at t (* carica X,Y,ZH,ZL:mTest_RS08_RAM *)
+    (load_from_source_at t (zero_memory t) (* carica source in ROM:mTest_RS08_prog *)
+     mTest_RS08_mult_source mTest_RS08_prog)
+    [ b1 ; b2 ; b3 ; b4 ] mTest_RS08_RAM)
+   (build_memory_type_of_mcu_version MC9RS08KA2 t)
+   (mk_byte8 x0 x0) (mk_byte8 x0 x0) (* non deterministici tutti a 0 *)
+   false false false false false false (* non deterministici tutti a 0 *)
+   ) mTest_RS08_prog)
+  true.
+
+(* parametrizzazione dell'enunciato del teorema mult *)
+lemma ok_mTest_RS08_mult_full_aux ≝
+ λt:memory_impl.
+ λX,Y:byte8.
+ match execute RS08 t
+       (TickOK ? (mTest_RS08_mult_status t X Y 〈x0,x0〉 〈x0,x0〉))
+       (12 +
+        (26 * (nat_of_byte8 Y)) +
+        (5 * (nat_of_byte8 (w16h (mul_b8 X Y))))) with
+  (* controllo che coincidano ALU,ZH,ZL *)
+  [ TickERR s _ ⇒ s
+  | TickSUSP s _ ⇒ s
+  (* FIXME: alla ALU azzero C perche' la funzione che ne
+            descrive il valore finale e' MOLTO complessa *)
+  | TickOK s ⇒ set_c_flag RS08 t s false ] =
+ (* alla fine per come e' scritto il programma: A=0 Z=true *)
+ set_pc_reg RS08 t (mTest_RS08_mult_status t X 〈x0,x0〉 (w16h (mul_b8 X Y)) (w16l (mul_b8 X Y)))
+  (mk_word16 〈x3,x8〉 〈x1,x6〉).
+
+(* dimostrazione senza svolgimento degli stati, immediata *)
+lemma ok_mTest_RS08_mult_full : 
+ ∀t:memory_impl.
+  ok_mTest_RS08_mult_full_aux t 〈xF,xF〉 〈x1,xE〉.
+ unfold ok_mTest_RS08_mult_full_aux;
+ intro;
+ elim t;
+ [ apply (eq_status_axiom mTest_RS08_RAM 3 RS08 MEM_FUNC); reflexivity ]
+ reflexivity.
+qed.
+
+(* ************************ *)
+(* TEST SU GRANULARITA' BIT *)
+(* ************************ *)
+
+definition mTest_bits_source : list byte8 ≝
+let m ≝ HCS08 in source_to_byte8 m (
+(* BEFORE: va a testare [0x0070]=0x00 *)
+(* [0x1860] 4clk *) (compile m ? MOV (maIMM1_to_DIR1 〈xF,xF〉 〈x7,x0〉) I)
+(* [0x1863] *)
+ ).
+
+(* creazione del processore+caricamento+impostazione registri *)
+definition mTest_bits_status ≝
+λt:memory_impl.
+λsub:oct.
+λb:byte8.
+ setweak_n_flag HCS08 t (* N<-1 *)
+  (set_pc_reg HCS08 t (* PC<-mTest_HCS08_prog *)
+   (start_of_mcu_version
+    MC9S08AW60 t
+    (load_from_source_at t
+     (load_from_source_at t (zero_memory t) (* carica b in RAM:mTest_HCS08_RAM *)
+       mTest_bits_source mTest_HCS08_prog) (* carica source in ROM:mTest_HCS08_prog *)
+      [ b ] mTest_HCS08_RAM)
+    (check_update_bit t (* setta mTest_HCS08_RAM,o come ROM *)
+     (build_memory_type_of_mcu_version MC9S08AW60 t)
+     mTest_HCS08_RAM sub MEM_READ_ONLY)
+    (mk_byte8 x0 x0) (mk_byte8 x0 x0) (* non deterministici tutti a 0 *)
+    false false false false false false) (* non deterministici tutti a 0 *)
+   mTest_HCS08_prog)
+  true.
+
+(* dimostrazione senza svolgimento degli stati, immediata *)
+lemma ok_mTest_bits_MEM_BITS_full :
+ ∀sub:oct.
+ execute HCS08 MEM_BITS (TickOK ? (mTest_bits_status MEM_BITS sub 〈x0,x0〉)) 4 =
+ TickOK ? (set_pc_reg HCS08 MEM_BITS (mTest_bits_status MEM_BITS sub (byte8_of_bits (setn_array8T sub ? (bits_of_byte8 〈xF,xF〉) false))) (* nuovo PC *)
+           (mk_word16 〈x1,x8〉 〈x6,x3〉)).
+ intro;
+ elim sub;
+ reflexivity.
+qed.
+
+lemma ok_mTest_bits_NO_MEM_BITS_full :
+ ∀sub:oct.
+ (match execute HCS08 MEM_FUNC (TickOK ? (mTest_bits_status MEM_FUNC sub 〈x0,x0〉)) 4 with
+  [ TickERR s _ ⇒ s | TickSUSP s _ ⇒ s | TickOK s ⇒ s ] =
+ set_pc_reg HCS08 MEM_FUNC (mTest_bits_status MEM_FUNC sub 〈xF,xF〉) (* nuovo PC *)
+           (mk_word16 〈x1,x8〉 〈x6,x3〉)) ∧
+ (match execute HCS08 MEM_TREE (TickOK ? (mTest_bits_status MEM_TREE sub 〈x0,x0〉)) 4 with
+  [ TickERR s _ ⇒ s | TickSUSP s _ ⇒ s | TickOK s ⇒ s ] =
+ set_pc_reg HCS08 MEM_TREE (mTest_bits_status MEM_TREE sub 〈xF,xF〉) (* nuovo PC *)
+           (mk_word16 〈x1,x8〉 〈x6,x3〉)).
+ intro;
+ elim sub;
+ split;
+ [ 1,3,5,7,9,11,13,15:
+   apply (eq_status_axiom mTest_HCS08_RAM 0 HCS08 MEM_FUNC); reflexivity; ]
+   apply (eq_status_axiom mTest_HCS08_RAM 0 HCS08 MEM_TREE); reflexivity;
+qed.
+
+(* svolgimento degli stati passo passo: esempio 
+
+ letin BEFORE ≝ ...
+
+ letin AFTER_ALU1 ≝ (match execute RS08 (TickOK ? BEFORE) 3 with
+  [ TickERR _ _ ⇒ get_alu RS08 BEFORE 
+  | TickSUSP _ _ ⇒ get_alu RS08 BEFORE 
+  | TickOK s ⇒ get_alu RS08 s ]);
+ normalize in AFTER_ALU1:(%);
+
+ letin AFTER_ALU2 ≝ (match execute RS08 (TickOK ? 
+                     (set_alu RS08 BEFORE AFTER_ALU1)) 3 with
+  [ TickERR _ _ ⇒ get_alu RS08 BEFORE 
+  | TickSUSP _ _ ⇒ get_alu RS08 BEFORE 
+  | TickOK s ⇒ get_alu RS08 s ]);
+ normalize in AFTER_ALU2:(%); 
+ clearbody AFTER_ALU1;
+
+ ...
+
+ letin AFTER_ALU6 ≝ (match execute RS08 (TickOK ? 
+                     (set_alu RS08 BEFORE AFTER_ALU5)) 5 with
+  [ TickERR _ _ ⇒ get_alu RS08 BEFORE 
+  | TickSUSP _ _ ⇒ get_alu RS08 BEFORE 
+  | TickOK s ⇒ get_alu RS08 s ]);
+ letin Y6 ≝ (match execute RS08 (TickOK ? 
+             (set_alu RS08 BEFORE AFTER_ALU5)) 5
+  return λ_.byte8 with
+  [ TickERR _ _ ⇒ (get_mem_desc RS08 BEFORE) 〈〈x0,x0〉:〈x2,x1〉〉
+  | TickSUSP _ _ ⇒ (get_mem_desc RS08 BEFORE) 〈〈x0,x0〉:〈x2,x1〉〉
+  | TickOK s ⇒ (get_mem_desc RS08 s) 〈〈x0,x0〉:〈x2,x1〉〉 ]);
+ normalize in AFTER_ALU6:(%); 
+ normalize in Y6:(%); 
+ clearbody AFTER_ALU5;
+
+ letin AFTER_ALU7 ≝ (match execute RS08 (TickOK ? 
+                     (set_mem_desc RS08 
+                      (set_alu RS08 BEFORE AFTER_ALU6)
+                       (load_from_source_at (get_mem_desc RS08 BEFORE)
+                        (mTest_RS08_mult_load 〈xF,xF〉 Y6 〈x0,x0〉 〈x0,x0〉) mTest_RS08_RAM) 
+                     )) 3 with
+  [ TickERR _ _ ⇒ get_alu RS08 BEFORE 
+  | TickSUSP _ _ ⇒ get_alu RS08 BEFORE 
+  | TickOK s ⇒ get_alu RS08 s ]);
+ normalize in AFTER_ALU7:(%); 
+ clearbody AFTER_ALU6;
+
+ ...
+*)
diff --git a/helm/software/matita/library/freescale/model.ma b/helm/software/matita/library/freescale/model.ma
new file mode 100644 (file)
index 0000000..56cd09a
--- /dev/null
@@ -0,0 +1,637 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/model/".
+
+(*include "/media/VIRTUOSO/freescale/status.ma".*)
+include "freescale/status.ma".
+
+(* *********************************** *)
+(* IMPOSTAZIONI SPECIFICHE DEI MODELLI *)
+(* *********************************** *)
+
+(* modelli di HC05 *)
+inductive HC05_mcu_model : Type ≝
+  MC68HC05J5A: HC05_mcu_model
+  (*..*).   
+
+(* modelli di HC08 *)
+inductive HC08_mcu_model : Type ≝
+  MC68HC08AB16A: HC08_mcu_model
+  (*..*). 
+
+(* modelli di HCS08 *)
+inductive HCS08_mcu_model : Type ≝
+  MC9S08AW60  : HCS08_mcu_model
+| MC9S08GB60  : HCS08_mcu_model
+| MC9S08GT60  : HCS08_mcu_model
+| MC9S08GB32  : HCS08_mcu_model
+| MC9S08GT32  : HCS08_mcu_model
+| MC9S08GT16  : HCS08_mcu_model
+| MC9S08GB60A : HCS08_mcu_model
+| MC9S08GT60A : HCS08_mcu_model
+| MC9S08GB32A : HCS08_mcu_model
+| MC9S08GT32A : HCS08_mcu_model
+| MC9S08GT16A : HCS08_mcu_model
+| MC9S08GT8A  : HCS08_mcu_model
+| MC9S08LC60  : HCS08_mcu_model
+| MC9S08LC36  : HCS08_mcu_model
+| MC9S08QD4   : HCS08_mcu_model
+| MC9S08QD2   : HCS08_mcu_model
+| MC9S08QG8   : HCS08_mcu_model
+| MC9S08QG4   : HCS08_mcu_model
+| MC9S08RC60  : HCS08_mcu_model
+| MC9S08RD60  : HCS08_mcu_model
+| MC9S08RE60  : HCS08_mcu_model
+| MC9S08RG60  : HCS08_mcu_model
+| MC9S08RC32  : HCS08_mcu_model
+| MC9S08RD32  : HCS08_mcu_model
+| MC9S08RE32  : HCS08_mcu_model
+| MC9S08RG32  : HCS08_mcu_model
+| MC9S08RC16  : HCS08_mcu_model
+| MC9S08RD16  : HCS08_mcu_model
+| MC9S08RE16  : HCS08_mcu_model
+| MC9S08RC8   : HCS08_mcu_model
+| MC9S08RD8   : HCS08_mcu_model
+| MC9S08RE8   : HCS08_mcu_model.
+
+(* modelli di RS08 *)
+inductive RS08_mcu_model : Type ≝
+  MC9RS08KA1 : RS08_mcu_model
+| MC9RS08KA2 : RS08_mcu_model.
+
+(* raggruppamento dei modelli *)
+inductive any_mcu_model : Type ≝
+  FamilyHC05  : HC05_mcu_model → any_mcu_model
+| FamilyHC08  : HC08_mcu_model → any_mcu_model
+| FamilyHCS08 : HCS08_mcu_model → any_mcu_model
+| FamilyRS08  : RS08_mcu_model → any_mcu_model.
+
+coercion cic:/matita/freescale/model/any_mcu_model.ind#xpointer(1/1/1).
+coercion cic:/matita/freescale/model/any_mcu_model.ind#xpointer(1/1/2).
+coercion cic:/matita/freescale/model/any_mcu_model.ind#xpointer(1/1/3).
+coercion cic:/matita/freescale/model/any_mcu_model.ind#xpointer(1/1/4).
+
+(* 
+condizioni errore interne alla MCU
+(Il programma viene caricato dal produttore in ROM o impostato in EEPROM)
+HC05: +illegal address during opcode fetch
+HC08: +illegal address duting opcode fetch (ILAD mascherabile)
+      +illegal opcode (ILOP mascherabile)
+
+(Il programma viene programmato nella FLASH)       
+HCS08: +illegal address during opcode fetch (ILAD mascherabile)
+       +illegal opcode (ILOP mascherabile)
+       +security = accesso a RAM e zone FLASH dichiarate sicure da zone sicure
+                   da' 0 in lettura e ignore in scrittura, [FLASH]SEC00:SEC01 (1,0)
+                   corrisponde a OFF, altre ON, disattivabile solo da modalita' sicura se
+                   opzione [FLASH]KEYEN e' 1 passando chiave a 8byte da modalita' sicura.
+                   Altrimenti disattivabile solo con FLASH mass erase. Obbiettivo
+                   e' impedire duplicazione di software di controllo, dato che esiste
+                   modalita' debugging. A restart viene ricaricata da FLASH impostazione
+                   sicurezza!
+RS08: +illegal address durting opcode fetch (ILAD) mascherabile
+      +illegal opcode (ILOP mascherabile)
+      +security = solo la FLASH e' considerata sicura. Stesso meccanismo di HCS08
+                  ma governato solo da [FLASH]SECD (0) OFF.Una volta attivato l'unica
+                  disattivazione e' la cancellazione della FLASH.
+*)
+
+(* memoria degli HC05 *)
+definition memory_type_of_FamilyHC05 ≝
+λm:HC05_mcu_model.match m with
+  [ MC68HC05J5A ⇒
+   [
+  (* astraggo molto *)
+  (* 0x0000-0x001F,0x0FF0: sarebbe memory mapped IO *)
+   
+     tripleT ??? 〈〈x0,x0〉:〈x8,x0〉〉 〈〈x0,x0〉:〈xF,xF〉〉 MEM_READ_WRITE (* 128B RAM+STACK *)
+   ; tripleT ??? 〈〈x0,x3〉:〈x0,x0〉〉 〈〈x0,xC〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 2560B USER ROM *)
+   ; tripleT ??? 〈〈x0,xE〉:〈x0,x0〉〉 〈〈x0,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 512B INTERNAL ROM *)
+    ]
+   (*..*)
+  ].
+
+(* memoria degli HC08 *)
+definition memory_type_of_FamilyHC08 ≝
+λm:HC08_mcu_model.match m with
+  [ MC68HC08AB16A ⇒
+   [
+  (* astraggo molto *) 
+  (* 0x0000-0x004F,0xFE00-0xFE01,0xFE03,
+     0xFE0C-0xFE11,0xFE1A-0xFE1D,0xFE1F: sarebbe memory mapped IO *)
+  (* 0x0500-0x057F,0xFE02,0xFE04-0xFE07,
+     0xFE09-0xFE0B,0xFE12-0xFE19,0xFE1E,0xFFC0-0xFFCF : sarebbe reserved *)
+      
+     tripleT ??? 〈〈x0,x0〉:〈x5,x0〉〉 〈〈x0,x2〉:〈x4,xF〉〉 MEM_READ_WRITE (* 512B RAM *)
+   ; tripleT ??? 〈〈x0,x8〉:〈x0,x0〉〉 〈〈x0,x9〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 512B EEPROM *)
+   ; tripleT ??? 〈〈xB,xE〉:〈x0,x0〉〉 〈〈xF,xD〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 16384B ROM *)
+   ; tripleT ??? 〈〈xF,xE〉:〈x2,x0〉〉 〈〈xF,xF〉:〈x5,x2〉〉 MEM_READ_ONLY  (* 307B ROM *)
+   ; tripleT ??? 〈〈xF,xF〉:〈xD,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 48B ROM *) ]
+   (*..*)
+  ].
+
+(* memoria degli HCS08 *)
+definition memory_type_of_FamilyHCS08 ≝
+λm:HCS08_mcu_model.match m with
+  [ MC9S08AW60 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x006F,0x1800-0x185F: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x7,x0〉〉 〈〈x0,x8〉:〈x6,xF〉〉 MEM_READ_WRITE (* 2048B RAM *)
+   ; tripleT ??? 〈〈x0,x8〉:〈x7,x0〉〉 〈〈x1,x7〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 3984B FLASH *)
+   ; tripleT ??? 〈〈x1,x8〉:〈x6,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 59296B FLASH *) ]
+  | MC9S08GB60 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x007F,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x8,x0〉〉 〈〈x1,x0〉:〈x7,xF〉〉 MEM_READ_WRITE (* 4096B RAM *)
+   ; tripleT ??? 〈〈x1,x0〉:〈x8,x0〉〉 〈〈x1,x7〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 1920B FLASH *)
+   ; tripleT ??? 〈〈x1,x8〉:〈x2,xC〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 59348B FLASH *) ]
+  | MC9S08GT60 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x007F,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x8,x0〉〉 〈〈x1,x0〉:〈x7,xF〉〉 MEM_READ_WRITE (* 4096B RAM *)
+   ; tripleT ??? 〈〈x1,x0〉:〈x8,x0〉〉 〈〈x1,x7〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 1920B FLASH *)
+   ; tripleT ??? 〈〈x1,x8〉:〈x2,xC〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 59348B FLASH *) ]
+  | MC9S08GB32 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x007F,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x8,x0〉〉 〈〈x0,x8〉:〈x7,xF〉〉 MEM_READ_WRITE (* 2048B RAM *)
+   ; tripleT ??? 〈〈x8,x0〉:〈x0,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 32768B FLASH *) ]
+  | MC9S08GT32 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x007F,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x8,x0〉〉 〈〈x0,x8〉:〈x7,xF〉〉 MEM_READ_WRITE (* 2048B RAM *)
+   ; tripleT ??? 〈〈x8,x0〉:〈x0,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 32768B FLASH *) ]
+  | MC9S08GT16 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x007F,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x8,x0〉〉 〈〈x0,x4〉:〈x7,xF〉〉 MEM_READ_WRITE (* 1024B RAM *)
+   ; tripleT ??? 〈〈xC,x0〉:〈x0,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 16384B FLASH *) ]
+  | MC9S08GB60A ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x007F,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x8,x0〉〉 〈〈x1,x0〉:〈x7,xF〉〉 MEM_READ_WRITE (* 4096B RAM *)
+   ; tripleT ??? 〈〈x1,x0〉:〈x8,x0〉〉 〈〈x1,x7〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 1920B FLASH *)
+   ; tripleT ??? 〈〈x1,x8〉:〈x2,xC〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 59348B FLASH *) ]
+  | MC9S08GT60A ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x007F,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x8,x0〉〉 〈〈x1,x0〉:〈x7,xF〉〉 MEM_READ_WRITE (* 4096B RAM *)
+   ; tripleT ??? 〈〈x1,x0〉:〈x8,x0〉〉 〈〈x1,x7〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 1920B FLASH *)
+   ; tripleT ??? 〈〈x1,x8〉:〈x2,xC〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 59348B FLASH *) ]
+  | MC9S08GB32A ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x007F,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x8,x0〉〉 〈〈x0,x8〉:〈x7,xF〉〉 MEM_READ_WRITE (* 2048B RAM *)
+   ; tripleT ??? 〈〈x8,x0〉:〈x0,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 32768B FLASH *) ]
+  | MC9S08GT32A ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x007F,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x8,x0〉〉 〈〈x0,x8〉:〈x7,xF〉〉 MEM_READ_WRITE (* 2048B RAM *)
+   ; tripleT ??? 〈〈x8,x0〉:〈x0,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 32768B FLASH *) ]
+  | MC9S08GT16A ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x007F,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x8,x0〉〉 〈〈x0,x8〉:〈x7,xF〉〉 MEM_READ_WRITE (* 2048B RAM *)
+   ; tripleT ??? 〈〈xC,x0〉:〈x0,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 16384B FLASH *) ]
+  | MC9S08GT8A ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x007F,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x8,x0〉〉 〈〈x0,x8〉:〈x7,xF〉〉 MEM_READ_WRITE (* 2048B RAM *)
+   ; tripleT ??? 〈〈xC,x0〉:〈x0,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 16384B FLASH *) ]
+  | MC9S08LC60 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x005F,0x1800-0x186F: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x6,x0〉〉 〈〈x1,x0〉:〈x5,xF〉〉 MEM_READ_WRITE (* 4096B RAM *)
+   ; tripleT ??? 〈〈x1,x0〉:〈x6,x0〉〉 〈〈x1,x7〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 1952B FLASH *)
+   ; tripleT ??? 〈〈x1,x8〉:〈x7,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 59280B FLASH *) ]
+  | MC9S08LC36 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x005F,0x1800-0x186F: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x6,x0〉〉 〈〈x0,xA〉:〈x5,xF〉〉 MEM_READ_WRITE (* 2560B RAM *)
+   ; tripleT ??? 〈〈x1,x8〉:〈x7,x0〉〉 〈〈x4,x8〉:〈x6,xF〉〉 MEM_READ_ONLY  (* 12288B FLASH *)
+   ; tripleT ??? 〈〈xA,x0〉:〈x0,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 24576B FLASH *) ]
+  | MC9S08QD4 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x005F,0x1800-0x184F: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x6,x0〉〉 〈〈x0,x1〉:〈x5,xF〉〉 MEM_READ_WRITE (* 256B RAM *)
+   ; tripleT ??? 〈〈xF,x0〉:〈x0,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 4096B FLASH *) ]
+  | MC9S08QD2 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x005F,0x1800-0x184F: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x6,x0〉〉 〈〈x0,x1〉:〈x5,xF〉〉 MEM_READ_WRITE (* 256B RAM *)
+   ; tripleT ??? 〈〈xF,x0〉:〈x0,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 4096B FLASH *) ]
+  | MC9S08QG8 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x005F,0x1800-0x184F: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x6,x0〉〉 〈〈x0,x2〉:〈x5,xF〉〉 MEM_READ_WRITE (* 512B RAM *)
+   ; tripleT ??? 〈〈xE,x0〉:〈x0,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 8192B FLASH *) ]
+  | MC9S08QG4 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x005F,0x1800-0x184F: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x6,x0〉〉 〈〈x0,x2〉:〈x5,xF〉〉 MEM_READ_WRITE (* 512B RAM *)
+   ; tripleT ??? 〈〈xE,x0〉:〈x0,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 8192B FLASH *) ]
+  | MC9S08RC60 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x0045,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x4,x6〉〉 〈〈x0,x8〉:〈x4,x5〉〉 MEM_READ_WRITE (* 2048B RAM *)
+   ; tripleT ??? 〈〈x0,x8〉:〈x4,x6〉〉 〈〈x1,x7〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 4026B FLASH *)
+   ; tripleT ??? 〈〈x1,x8〉:〈x2,xC〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 59348B FLASH *) ]
+  | MC9S08RD60 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x0045,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x4,x6〉〉 〈〈x0,x8〉:〈x4,x5〉〉 MEM_READ_WRITE (* 2048B RAM *)
+   ; tripleT ??? 〈〈x0,x8〉:〈x4,x6〉〉 〈〈x1,x7〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 4026B FLASH *)
+   ; tripleT ??? 〈〈x1,x8〉:〈x2,xC〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 59348B FLASH *) ]
+  | MC9S08RE60 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x0045,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x4,x6〉〉 〈〈x0,x8〉:〈x4,x5〉〉 MEM_READ_WRITE (* 2048B RAM *)
+   ; tripleT ??? 〈〈x0,x8〉:〈x4,x6〉〉 〈〈x1,x7〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 4026B FLASH *)
+   ; tripleT ??? 〈〈x1,x8〉:〈x2,xC〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 59348B FLASH *) ]
+  | MC9S08RG60 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x0045,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x4,x6〉〉 〈〈x0,x8〉:〈x4,x5〉〉 MEM_READ_WRITE (* 2048B RAM *)
+   ; tripleT ??? 〈〈x0,x8〉:〈x4,x6〉〉 〈〈x1,x7〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 4026B FLASH *)
+   ; tripleT ??? 〈〈x1,x8〉:〈x2,xC〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 59348B FLASH *) ]
+  | MC9S08RC32 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x0045,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x4,x6〉〉 〈〈x0,x8〉:〈x4,x5〉〉 MEM_READ_WRITE (* 2048B RAM *)
+   ; tripleT ??? 〈〈x8,x0〉:〈x0,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 32768B FLASH *) ]
+  | MC9S08RD32 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x0045,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x4,x6〉〉 〈〈x0,x8〉:〈x4,x5〉〉 MEM_READ_WRITE (* 2048B RAM *)
+   ; tripleT ??? 〈〈x8,x0〉:〈x0,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 32768B FLASH *) ]
+  | MC9S08RE32 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x0045,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x4,x6〉〉 〈〈x0,x8〉:〈x4,x5〉〉 MEM_READ_WRITE (* 2048B RAM *)
+   ; tripleT ??? 〈〈x8,x0〉:〈x0,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 32768B FLASH *) ]
+  | MC9S08RG32 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x0045,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x4,x6〉〉 〈〈x0,x8〉:〈x4,x5〉〉 MEM_READ_WRITE (* 2048B RAM *)
+   ; tripleT ??? 〈〈x8,x0〉:〈x0,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 32768B FLASH *) ]
+  | MC9S08RC16 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x003F,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x4,x0〉〉 〈〈x0,x4〉:〈x3,xF〉〉 MEM_READ_WRITE (* 1024B RAM *)
+   ; tripleT ??? 〈〈xC,x0〉:〈x0,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 16384B FLASH *) ]
+  | MC9S08RD16 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x003F,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x4,x0〉〉 〈〈x0,x4〉:〈x3,xF〉〉 MEM_READ_WRITE (* 1024B RAM *)
+   ; tripleT ??? 〈〈xC,x0〉:〈x0,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 16384B FLASH *) ]
+  | MC9S08RE16 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x003F,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x4,x0〉〉 〈〈x0,x4〉:〈x3,xF〉〉 MEM_READ_WRITE (* 1024B RAM *)
+   ; tripleT ??? 〈〈xC,x0〉:〈x0,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 16384B FLASH *) ]
+  | MC9S08RC8 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x003F,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x4,x0〉〉 〈〈x0,x4〉:〈x3,xF〉〉 MEM_READ_WRITE (* 1024B RAM *)
+   ; tripleT ??? 〈〈xE,x0〉:〈x0,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 8192B FLASH *) ]
+  | MC9S08RD8 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x003F,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x4,x0〉〉 〈〈x0,x4〉:〈x3,xF〉〉 MEM_READ_WRITE (* 1024B RAM *)
+   ; tripleT ??? 〈〈xE,x0〉:〈x0,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 8192B FLASH *) ]
+  | MC9S08RE8 ⇒
+   [ 
+  (* astraggo molto *)
+  (* 0x0000-0x003F,0x1800-0x182B: sarebbe memory mapped IO *)
+  
+     tripleT ??? 〈〈x0,x0〉:〈x4,x0〉〉 〈〈x0,x4〉:〈x3,xF〉〉 MEM_READ_WRITE (* 1024B RAM *)
+   ; tripleT ??? 〈〈xE,x0〉:〈x0,x0〉〉 〈〈xF,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 8192B FLASH *) ]
+  ].
+
+(* memoria dei RS08 *)
+definition memory_type_of_FamilyRS08 ≝
+λm:RS08_mcu_model.match m with
+  [ MC9RS08KA1 ⇒
+   [
+     tripleT ??? 〈〈x0,x0〉:〈x0,x0〉〉 〈〈x0,x0〉:〈x0,xE〉〉 MEM_READ_WRITE (* 15B RAM *)
+  (* [000F] e' il registro X *)
+  (* 0x0010-0x001E sarebbe memory mapped IO, proviamo per completezza con RAM *)
+   ; tripleT ??? 〈〈x0,x0〉:〈x1,x0〉〉 〈〈x0,x0〉:〈x1,xE〉〉 MEM_READ_WRITE (* 15B MEMORY MAPPED IO *)
+  (* [001F] e' il registro PAGESEL *)
+   ; tripleT ??? 〈〈x0,x0〉:〈x2,x0〉〉 〈〈x0,x0〉:〈x4,xF〉〉 MEM_READ_WRITE (* 48B RAM *)
+  (* [00C0-00FF] mappato da PAGESEL su [00pp pppp ppxx xxxx] *)
+   ; tripleT ??? 〈〈x0,x0〉:〈xC,x0〉〉 〈〈x0,x0〉:〈xF,xF〉〉 MEM_READ_WRITE (* 64B RAM PAGING *)
+  (* 0x0200-0x023F sarebbe memory mapped IO, proviamo per completezza con RAM *)
+   ; tripleT ??? 〈〈x0,x2〉:〈x0,x0〉〉 〈〈x0,x2〉:〈x3,xF〉〉 MEM_READ_WRITE (* 64B MEMORY MAPPED IO *)
+   ; tripleT ??? 〈〈x3,xC〉:〈x0,x0〉〉 〈〈x3,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 1024B FLASH *) ]
+  | MC9RS08KA2 ⇒
+   [ 
+     tripleT ??? 〈〈x0,x0〉:〈x0,x0〉〉 〈〈x0,x0〉:〈x0,xE〉〉 MEM_READ_WRITE (* 15B RAM *)
+  (* [000F] e' il registro X *)
+  (* 0x0010-0x001E sarebbe memory mapped IO, proviamo per completezza con RAM *)
+   ; tripleT ??? 〈〈x0,x0〉:〈x1,x0〉〉 〈〈x0,x0〉:〈x1,xE〉〉 MEM_READ_WRITE (* 15B MEMORY MAPPED IO *)
+  (* [001F] e' il registro PAGESEL *)
+   ; tripleT ??? 〈〈x0,x0〉:〈x2,x0〉〉 〈〈x0,x0〉:〈x4,xF〉〉 MEM_READ_WRITE (* 48B RAM *)
+  (* [00C0-00FF] mappato da PAGESEL su [00pp pppp ppxx xxxx] *)
+   ; tripleT ??? 〈〈x0,x0〉:〈xC,x0〉〉 〈〈x0,x0〉:〈xF,xF〉〉 MEM_READ_WRITE (* 64B RAM PAGING *)
+  (* 0x0200-0x023F sarebbe memory mapped IO, proviamo per completezza con RAM *)
+   ; tripleT ??? 〈〈x0,x2〉:〈x0,x0〉〉 〈〈x0,x2〉:〈x3,xF〉〉 MEM_READ_WRITE (* 64B MEMORY MAPPED IO *)
+   ; tripleT ??? 〈〈x3,x8〉:〈x0,x0〉〉 〈〈x3,xF〉:〈xF,xF〉〉 MEM_READ_ONLY  (* 2048B FLASH *) ]
+  ].
+
+(* ∀modello.descrizione della memoria installata *)
+definition memory_type_of_mcu_version ≝
+λmcu:any_mcu_model.match mcu with
+ [ FamilyHC05  m ⇒ memory_type_of_FamilyHC05 m
+ | FamilyHC08  m ⇒ memory_type_of_FamilyHC08 m
+ | FamilyHCS08 m ⇒ memory_type_of_FamilyHCS08 m
+ | FamilyRS08  m ⇒ memory_type_of_FamilyRS08 m
+ ].
+
+(* dato un modello costruisce un descrittore a partire dalla lista precedente *)
+definition build_memory_type_of_mcu_version ≝
+λmcu:any_mcu_model.λt:memory_impl.
+ let rec aux param result ≝
+  match param with
+   [ nil ⇒ result
+   | cons hd tl ⇒ 
+    aux tl (check_update_ranged t result (fst3T ??? hd) (snd3T ??? hd) (thd3T ??? hd)) ]
+ in aux (memory_type_of_mcu_version mcu) (out_of_bound_memory t).
+
+(* sarebbe programma da caricare/zero_memory, ora test *)
+definition memory_of_mcu_version ≝
+λmcu:any_mcu_model.λt:memory_impl.match mcu with
+ [ FamilyHC05 m ⇒ match m with
+  [ MC68HC05J5A ⇒ zero_memory t
+    (*..*)
+  ]
+ | FamilyHC08 m ⇒ match m with
+  [ MC68HC08AB16A ⇒ zero_memory t
+    (*..*)
+  ]
+ (* tralascio l'enumerazione dei casi, per ora e' tutto 0 *)
+ | FamilyHCS08 _ ⇒ zero_memory t
+ | FamilyRS08 _ ⇒ zero_memory t
+ ].
+
+(* inferire la mcu a partire dal modello *)
+definition mcu_of_model ≝
+λm:any_mcu_model.match m with
+ [ FamilyHC05  _ ⇒ HC05
+ | FamilyHC08  _ ⇒ HC08
+ | FamilyHCS08 _ ⇒ HCS08
+ | FamilyRS08  _ ⇒ RS08
+ ].
+
+(* 
+   parametrizzati i non deterministici rispetto a tutti i valori casuali
+   che verranno dati dall'esterno di tipo byte8 (ndby1-2) e bool (ndbo1-5).
+   l'ACCENSIONE e' totalmente equivalente ad un reset causato da calo di tensione
+   (reset V-low), la memoria ed il check puo' essere passata, per esempio da
+   - (memory_of_mcu_version MC68HC05J5A)
+   - (build_memory_type_of_mcu_version MC68HC05J5A)
+*)
+definition start_of_mcu_version ≝
+λmcu:any_mcu_model.λt:memory_impl.
+let pc_reset_h ≝ 〈〈xF,xF〉:〈xF,xE〉〉 in
+let pc_reset_l ≝ 〈〈xF,xF〉:〈xF,xF〉〉 in
+let pc_RS08_reset ≝ 〈〈xF,xF〉:〈xF,xD〉〉 in
+let sp_reset ≝ 〈〈x0,x0〉:〈xF,xF〉〉 in
+match mcu
+ return λmcu:any_mcu_model.(aux_mem_type t) → (aux_chk_type t) → byte8 → byte8 →
+        bool → bool → bool → bool → bool → bool → any_status (mcu_of_model mcu) t
+with
+(* HC05: parzialmente non deterministico *)
+ [ FamilyHC05 m ⇒ λmem:aux_mem_type t.λchk:aux_chk_type t.
+                  λndby1,ndby2:byte8.λirqfl,ndbo1,ndbo2,ndbo3,ndbo4,ndbo5:bool.
+  let build ≝ λspm,spf,pcm.
+  let fetched_pc ≝ mk_word16 (mem_read_abs t mem (and_w16 pc_reset_h pcm)) 
+                             (mem_read_abs t mem (and_w16 pc_reset_l pcm)) in
+   mk_any_status HC05 t
+    (mk_alu_HC05
+     (* acc_low: ? *) ndby1 (* indx_low: ? *) ndby2 
+     (* sp: reset *) (or_w16 (and_w16 sp_reset spm) spf) spm spf
+     (* pc: reset+fetch *) (and_w16 fetched_pc pcm) pcm
+     (* H: ? *) ndbo1 (* I: reset *) true (* N: ? *) ndbo2
+     (* Z: ? *) ndbo3 (* C: ? *) ndbo4 (* IRQ: ? *) irqfl)
+     (* mem *) mem
+     (* chk *) chk
+     (* clk: reset *) (None ?) in
+  match m with
+   [ MC68HC05J5A ⇒ build 〈〈x0,x0〉:〈x3,xF〉〉 〈〈x0,x0〉:〈xC,x0〉〉 〈〈x0,xF〉:〈xF,xF〉〉
+     (*..*)
+   ]
+
+(* HC08: parzialmente non deterministico *)
+ | FamilyHC08 m ⇒ λmem:aux_mem_type t.λchk:aux_chk_type t.
+                  λndby1,ndby2:byte8.λirqfl,ndbo1,ndbo2,ndbo3,ndbo4,ndbo5:bool.
+  let build ≝
+  let fetched_pc ≝ mk_word16 (mem_read_abs t mem pc_reset_h) 
+                             (mem_read_abs t mem pc_reset_l) in
+   mk_any_status HC08 t
+    (mk_alu_HC08
+     (* acc_low: ? *) ndby1 (* indw_low: ? *) ndby2 (* indx_high: reset *)  〈x0,x0〉
+     (* sp: reset *) sp_reset (* pc: reset+fetch *) fetched_pc
+     (* V: ? *) ndbo1 (* H: ? *) ndbo2 (* I: reset *) true
+     (* N: ? *) ndbo3 (* Z: ? *) ndbo4 (* C: ? *) ndbo5 (* IRQ: ? *) irqfl)
+     (* mem *) mem
+     (* chk *) chk
+     (* clk: reset *) (None ?) in
+  (* tralascio l'enumerazione dei casi, tanto non ci sono maschere da impostare *)
+  build
+
+(* HCS08: parzialmente non deterministico *)
+ | FamilyHCS08 m ⇒ λmem:aux_mem_type t.λchk:aux_chk_type t.
+                   λndby1,ndby2:byte8.λirqfl,ndbo1,ndbo2,ndbo3,ndbo4,ndbo5:bool.
+  let build ≝
+  let fetched_pc ≝ mk_word16 (mem_read_abs t mem pc_reset_h)
+                             (mem_read_abs t mem pc_reset_l) in
+   mk_any_status HCS08 t
+    (mk_alu_HC08
+     (* acc_low: ? *) ndby1 (* indw_low: ? *) ndby2 (* indx_high: reset *)  〈x0,x0〉
+     (* sp: reset *) sp_reset (* pc: reset+fetch *) fetched_pc
+     (* V: ? *) ndbo1 (* H: ? *) ndbo2 (* I: reset *) true
+     (* N: ? *) ndbo3 (* Z: ? *) ndbo4 (* C: ? *) ndbo5 (* IRQ: ? *) irqfl)
+     (* mem *) mem
+     (* chk *) chk
+     (* clk: reset *) (None ?) in
+  (* tralascio l'enumerazione dei casi, tanto non ci sono maschere da impostare *)
+  build
+
+(* RS08: deterministico *)
+ | FamilyRS08 m ⇒ λmem:aux_mem_type t.λchk:aux_chk_type t.
+                  λndby1,ndby2:byte8.λirqfl,ndbo1,ndbo2,ndbo3,ndbo4,ndbo5:bool.
+  let build ≝ λpcm.
+   mk_any_status RS08 t
+    (mk_alu_RS08
+     (* acc_low: reset *)  〈x0,x0〉
+     (* pc: reset *) (and_w16 pc_RS08_reset pcm) pcm
+     (* spc: reset *) (and_w16 pc_RS08_reset pcm)
+     (* xm: reset *) 〈x0,x0〉 (* psm: *) 〈x8,x0〉
+     (* Z: reset *) false (* C: reset *) false)
+     (* mem *) mem
+     (* chk *) chk
+     (* clk: reset *) (None ?) in
+  (* tralascio l'enumerazione dei casi, tanto i valori sono uguali *)
+  build 〈〈x3,xF〉:〈xF,xF〉〉
+ ].
+
+(* 
+   cio' che non viene resettato mantiene il valore precedente: nella documentazione
+   viene riportato come "unaffected"/"indeterminate"/"unpredictable"
+   il soft RESET e' diverso da un calo di tensione e la ram non variera'
+*)
+definition reset_of_mcu ≝
+λm:mcu_type.λt:memory_impl.
+let pc_reset_h ≝ 〈〈xF,xF〉:〈xF,xE〉〉 in
+let pc_reset_l ≝ 〈〈xF,xF〉:〈xF,xF〉〉 in
+let pc_RS08_reset ≝ 〈〈xF,xF〉:〈xF,xD〉〉 in
+let sp_reset ≝ 〈〈x0,x0〉:〈xF,xF〉〉 in
+ match m return λm:mcu_type.(any_status m t) → (any_status m t) with
+(* HC05: parzialmente non deterministico *)
+  [ HC05 ⇒ λs:any_status HC05 t.match s with
+   [ mk_any_status alu mem chk clk ⇒ match alu with
+    [ mk_alu_HC05 acclow indxlow _ spm spf _ pcm hfl _ nfl zfl cfl irqfl ⇒
+    let fetched_pc ≝ mk_word16 (mem_read_abs t mem (and_w16 pc_reset_h pcm))
+                               (mem_read_abs t mem (and_w16 pc_reset_l pcm)) in
+    mk_any_status HC05 t
+     (mk_alu_HC05
+      (* acc_low: inv. *) acclow (* indx_low: inv. *) indxlow
+      (* sp: reset *) (or_w16 (and_w16 sp_reset spm) spf) spm spf
+      (* pc: reset+fetch *) (and_w16 fetched_pc pcm) pcm
+      (* H: inv. *) hfl (* I: reset *) true (* N: inv. *) nfl
+      (* Z: inv. *) zfl (* C: inv. *) cfl (* IRQ: inv *) irqfl)
+      (* mem: inv. *) mem
+      (* chk: inv. *) chk
+      (* clk: reset *) (None ?) ]]
+(* HC08: parzialmente non deterministico *)
+  | HC08 ⇒ λs:any_status HC08 t.match s with
+   [ mk_any_status alu mem chk clk ⇒ match alu with
+    [ mk_alu_HC08 acclow indxlow _ _ _ vfl hfl _ nfl zfl cfl irqfl ⇒ 
+    let fetched_pc ≝ mk_word16 (mem_read_abs t mem pc_reset_h)
+                               (mem_read_abs t mem pc_reset_l) in
+    mk_any_status HC08 t
+     (mk_alu_HC08
+      (* acc_low: inv. *) acclow (* indx_low: inv. *) indxlow (* indx_high: reset *) 〈x0,x0〉
+      (* sp: reset *) sp_reset (* pc: reset+fetch *) fetched_pc
+      (* V: inv. *) vfl (* H: inv. *) hfl (* I: reset *) true
+      (* N: inv. *) nfl (* Z: inv. *) zfl (* C: inv. *) cfl (* IRQ: inv *) irqfl)
+      (* mem: inv. *) mem
+      (* chk: inv. *) chk
+      (* clk: reset *) (None ?) ]]
+(* HCS08: parzialmente non deterministico *)
+  | HCS08 ⇒ λs:any_status HCS08 t.match s with
+   [ mk_any_status alu mem chk clk ⇒ match alu with
+    [ mk_alu_HC08 acclow indxlow _ _ _ vfl hfl _ nfl zfl cfl irqfl ⇒ 
+    let fetched_pc ≝ mk_word16 (mem_read_abs t mem pc_reset_h)
+                               (mem_read_abs t mem pc_reset_l) in
+    mk_any_status HCS08 t
+     (mk_alu_HC08
+      (* acc_low: inv. *) acclow (* indx_low: inv. *) indxlow (* indx_high: reset *) 〈x0,x0〉
+      (* sp: reset *) sp_reset (* pc: reset+fetch *) fetched_pc
+      (* V: inv. *) vfl (* H: inv. *) hfl (* I: reset *) true
+      (* N: inv. *) nfl (* Z: inv. *) zfl (* C: inv. *) cfl (* IRQ: inv *) irqfl)
+      (* mem: inv. *) mem
+      (* chk: inv. *) chk
+      (* clk: reset *) (None ?) ]]
+(* RS08: deterministico *)
+  | RS08 ⇒ λs:any_status RS08 t.match s with
+   [ mk_any_status alu mem chk clk ⇒ match alu with
+    [ mk_alu_RS08 _ _ pcm _ _ _ _ _ ⇒
+    mk_any_status RS08 t
+     (mk_alu_RS08 
+      (* acc_low: reset *) 〈x0,x0〉
+      (* pc: reset *) (and_w16 pc_RS08_reset pcm) pcm
+      (* spc: reset *) (and_w16 pc_RS08_reset pcm)
+      (* xm: reset *) 〈x0,x0〉 (* psm: reset *) 〈x8,x0〉
+      (* Z: reset *) false (* C: reset *) false)
+      (* mem: inv. *) mem
+      (* chk: inv. *) chk
+      (* clk: reset *) (None ?) ]]
+  ].
diff --git a/helm/software/matita/library/freescale/multivm.ma b/helm/software/matita/library/freescale/multivm.ma
new file mode 100644 (file)
index 0000000..35e8328
--- /dev/null
@@ -0,0 +1,1336 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/multivm/".
+
+(*include "/media/VIRTUOSO/freescale/load_write.ma".*)
+include "freescale/load_write.ma".
+
+(* ************************************************ *)
+(* LOGICHE AUSILIARE CHE ACCOMUNANO PIU' OPERAZIONI *)
+(* ************************************************ *)
+
+(* NB: dentro il codice i commenti cercano di spiegare la logica
+       secondo quanto riportato dalle specifiche delle mcu *)
+
+(* NB: notare che tranne nei casi in cui PC viene modificato esplicitamente
+       il suo avanzamento viene delegato totalmente agli strati inferiori
+       I) avanzamento dovuto al decode degli op (fetch)
+       II) avanzamento dovuto al caricamento degli argomenti (multi_mode_load/write)
+       la modifica effettiva avviene poi centralizzata in tick *)
+
+(* A = [true] fAMC(A,M,C), [false] A *)
+(* cioe' in caso di false l'operazione viene eseguita ma modifica solo i flag *)
+(* fAMC e' la logica da applicare: somma con/senza carry *)
+definition execute_ADC_ADD_aux ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.λsetflag:bool.
+λfAMC:byte8 → byte8 → bool → ProdT byte8 bool.
+ opt_map ?? (multi_mode_load true m t s cur_pc i)
+  (λS_M_PC.match S_M_PC with
+   [ tripleT s_tmp1 M_op new_pc ⇒
+    let A_op ≝ get_acc_8_low_reg m t s_tmp1 in
+    match fAMC A_op M_op (get_c_flag m t s_tmp1) with
+     [ pairT R_op carry ⇒
+      let A7 ≝ MSB_b8 A_op in let M7 ≝ MSB_b8 M_op in let R7 ≝ MSB_b8 R_op in
+      let A3 ≝ MSB_ex (b8l A_op) in let M3 ≝ MSB_ex (b8l M_op) in let R3 ≝ MSB_ex (b8l R_op) in
+      (* A = [true] fAMC(A,M,C), [false] A *)
+      let s_tmp2 ≝ match setflag with [ true ⇒ set_acc_8_low_reg m t s_tmp1 R_op | false ⇒ s_tmp1 ] in
+      (* Z = nR7&nR6&nR5&nR4&nR3&nR2&nR1&nR0 *)
+      let s_tmp3 ≝ set_z_flag m t s_tmp2 (eq_b8 R_op 〈x0,x0〉) in
+      (* C = A7&M7 | M7&nR7 | nR7&A7 *)
+      let s_tmp4 ≝ set_c_flag m t s_tmp3 ((A7⊗M7) ⊕ (M7⊗(⊖R7)) ⊕ ((⊖R7)⊗A7)) in
+      (* N = R7 *)
+      let s_tmp5 ≝ setweak_n_flag m t s_tmp4 R7 in
+      (* H = A3&M3 | M3&nR3 | nR3&A3 *)
+      let s_tmp6 ≝ setweak_h_flag m t s_tmp5 ((A3⊗M3) ⊕ (M3⊗(⊖R3)) ⊕ ((⊖R3)⊗A3)) in
+      (* V = A7&M7&nR7 | nA7&nM7&R7 *)
+      let s_tmp7 ≝ setweak_v_flag m t s_tmp6 ((A7⊗M7⊗(⊖R7)) ⊕ ((⊖A7)⊗(⊖M7)⊗R7)) in
+      (* newpc = nextpc *)
+      Some ? (pairT ?? s_tmp7 new_pc) ]]).
+
+(* A = [true] fAM(A,M), [false] A *)
+(* cioe' in caso di false l'operazione viene eseguita ma modifica solo i flag *)
+(* fAM e' la logica da applicare: and/xor/or *)
+definition execute_AND_BIT_EOR_ORA_aux ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.λsetflag:bool.
+λfAM:byte8 → byte8 → byte8.
+ opt_map ?? (multi_mode_load true m t s cur_pc i)
+  (λS_M_PC.match S_M_PC with
+   [ tripleT s_tmp1 M_op new_pc ⇒
+    let R_op ≝ fAM (get_acc_8_low_reg m t s_tmp1) M_op in
+    (* A = [true] fAM(A,M), [false] A *) 
+    let s_tmp2 ≝ match setflag with [ true ⇒ set_acc_8_low_reg m t s_tmp1 R_op | false ⇒ s_tmp1 ] in
+    (* Z = nR7&nR6&nR5&nR4&nR3&nR2&nR1&nR0 *)
+    let s_tmp3 ≝ set_z_flag m t s_tmp2 (eq_b8 R_op 〈x0,x0〉) in
+    (* N = R7 *) 
+    let s_tmp4 ≝ setweak_n_flag m t s_tmp3 (MSB_b8 R_op) in
+    (* V = 0 *) 
+    let s_tmp5 ≝ setweak_v_flag m t s_tmp4 false in
+    (* newpc = nextpc *)
+    Some ? (pairT ?? s_tmp5 new_pc) ]).
+
+(* M = fMC(M,C) *)
+(* fMC e' la logica da applicare: rc_/ro_/sh_ *)
+definition execute_ASL_ASR_LSR_ROL_ROR_aux ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+λfMC:byte8 → bool → ProdT byte8 bool.
+ opt_map ?? (multi_mode_load true m t s cur_pc i)
+  (λS_M_PC.match S_M_PC with
+   [ tripleT s_tmp1 M_op _ ⇒
+    match fMC M_op (get_c_flag m t s_tmp1) with [ pairT R_op carry ⇒
+    (* M = fMC(M,C) *)
+    opt_map ?? (multi_mode_write true m t s_tmp1 cur_pc i R_op)
+     (λS_PC.match S_PC with
+      [ pairT s_tmp2 new_pc ⇒
+      (* C = carry *)
+      let s_tmp3 ≝ set_c_flag m t s_tmp2 carry in
+      (* Z = nR7&nR6&nR5&nR4&nR3&nR2&nR1&nR0 *)
+      let s_tmp4 ≝ set_z_flag m t s_tmp3 (eq_b8 R_op 〈x0,x0〉) in
+      (* N = R7 *)
+      let s_tmp5 ≝ setweak_n_flag m t s_tmp4 (MSB_b8 R_op) in
+      (* V = R7 ⊙ carry *)
+      let s_tmp6 ≝ setweak_v_flag m t s_tmp5 ((MSB_b8 R_op) ⊙ carry) in
+      (* newpc = nextpc *)
+      Some ? (pairT ?? s_tmp6 new_pc) ])]]).
+
+(* estensione del segno byte → word *)
+definition byte_extension ≝
+λb:byte8.〈match MSB_b8 b with [ true ⇒ 〈xF,xF〉 | false ⇒ 〈x0,x0〉 ]:b〉.
+
+(* branch con byte+estensione segno *)
+definition branched_pc ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λcur_pc:word16.λb:byte8.
+ get_pc_reg m t (set_pc_reg m t s (plus_w16nc cur_pc (byte_extension b))).
+
+(* if COND=1 branch *)
+(* tutti i branch calcoleranno la condizione e la passeranno qui *)
+definition execute_any_BRANCH ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.λfCOND:bool.
+ opt_map ?? (multi_mode_load true m t s cur_pc i)
+  (λS_M_PC.match S_M_PC with
+   [ tripleT s_tmp1 M_op new_pc ⇒
+    (* if true, branch *) 
+    match fCOND with
+     (* newpc = nextpc + rel *)
+     [ true ⇒ Some ? (pairT ?? s_tmp1 (branched_pc m t s_tmp1 new_pc M_op))
+     (* newpc = nextpc *)
+     | false ⇒ Some ? (pairT ?? s_tmp1 new_pc) ]]).
+
+(* Mn = filtered optval *) 
+(* il chiamante passa 0x00 per azzerare, 0xFF per impostare il bit di M *)
+definition execute_BCLRn_BSETn_aux ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.λoptval:byte8.
+ (* Mn = filtered optval *)
+ opt_map ?? (multi_mode_write true m t s cur_pc i optval)
+  (λS_PC.match S_PC with
+   (* newpc = nextpc *)
+   [ pairT s_tmp1 new_pc ⇒ Some ? (pairT ?? s_tmp1 new_pc) ]).
+
+(* if COND(Mn) branch *)
+(* il chiamante passa la logica da testare (0x00,¬0x00) e poi si salta *)
+definition execute_BRCLRn_BRSETn_aux ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.λfCOND:byte8 → bool.
+ opt_map ?? (multi_mode_load false m t s cur_pc i)
+  (λS_M_PC.match S_M_PC with
+   [ tripleT s_tmp1 M_op new_pc ⇒ match M_op with
+    [ mk_word16 MH_op ML_op ⇒
+     (* if COND(Mn) branch *)
+     match fCOND MH_op with
+      (* newpc = nextpc + rel *)
+      [ true ⇒ Some ? (pairT ?? s_tmp1 (branched_pc m t s_tmp1 new_pc ML_op))
+      (* newpc = nextpc *)
+      | false ⇒ Some ? (pairT ?? s_tmp1 new_pc) ]]]).
+
+(* M = fM(M) *)
+(* fM e' la logica da applicare: not/neg/++/-- *)
+definition execute_COM_DEC_INC_NEG_aux ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+λfM:byte8 → byte8.λfV:bool → bool → bool.λfC:bool → byte8 → bool.
+ opt_map ?? (multi_mode_load true m t s cur_pc i)
+  (λS_M_PC.match S_M_PC with
+   [ tripleT s_tmp1 M_op _ ⇒
+    let R_op ≝ fM M_op in
+    (* M = fM(M) *)
+    opt_map ?? (multi_mode_write true m t s_tmp1 cur_pc i R_op)
+     (λS_PC.match S_PC with
+      [ pairT s_tmp2 new_pc ⇒
+      (* C = fCR (C,R) *)
+      let s_tmp3 ≝ set_c_flag m t s_tmp2 (fC (get_c_flag m t s_tmp2) R_op) in
+      (* Z = nR7&nR6&nR5&nR4&nR3&nR2&nR1&nR0 *)
+      let s_tmp4 ≝ set_z_flag m t s_tmp3 (eq_b8 R_op 〈x0,x0〉) in
+      (* N = R7 *)
+      let s_tmp5 ≝ setweak_n_flag m t s_tmp4 (MSB_b8 R_op) in
+      (* V = fV (M7,R7) *)
+      let s_tmp6 ≝ setweak_v_flag m t s_tmp5 (fV (MSB_b8 M_op) (MSB_b8 R_op)) in
+      (* newpc = nextpc *)
+      Some ? (pairT ?? s_tmp6 new_pc) ])]).
+
+(* A = [true] fAMC(A,M,C), [false] A *)
+(* cioe' in caso di false l'operazione viene eseguita ma modifica solo i flag *)
+(* fAMC e' la logica da applicare: sottrazione con/senza carry *)
+definition execute_SBC_SUB_aux ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.λsetflag:bool.
+λfAMC:byte8 → byte8 → bool → ProdT byte8 bool.
+ opt_map ?? (multi_mode_load true m t s cur_pc i)
+  (λS_M_PC.match S_M_PC with
+   [ tripleT s_tmp1 M_op new_pc ⇒
+    let A_op ≝ get_acc_8_low_reg m t s_tmp1 in
+    match fAMC A_op M_op (get_c_flag m t s_tmp1) with
+     [ pairT R_op carry ⇒
+      let A7 ≝ MSB_b8 A_op in let M7 ≝ MSB_b8 M_op in let R7 ≝ MSB_b8 R_op in
+      (* A = [true] fAMC(A,M,C), [false] A *)
+      let s_tmp2 ≝ match setflag with [ true ⇒ set_acc_8_low_reg m t s_tmp1 R_op | false ⇒ s_tmp1 ] in
+      (* Z = nR7&nR6&nR5&nR4&nR3&nR2&nR1&nR0 *)
+      let s_tmp3 ≝ set_z_flag m t s_tmp2 (eq_b8 R_op 〈x0,x0〉) in
+      (* C = nA7&M7 | M7&R7 | R7&nA7 *)
+      let s_tmp4 ≝ set_c_flag m t s_tmp3 (((⊖A7)⊗M7) ⊕ (M7⊗R7) ⊕ (R7⊗(⊖A7))) in
+      (* N = R7 *) 
+      let s_tmp5 ≝ setweak_n_flag m t s_tmp4 R7 in
+      (* V = A7&nM7&nR7 | nA7&M7&R7 *)
+      let s_tmp6 ≝ setweak_v_flag m t s_tmp5 ((A7⊗(⊖M7)⊗(⊖R7)) ⊕ ((⊖A7)⊗M7⊗R7)) in
+      (* newpc = nextpc *)
+      Some ? (pairT ?? s_tmp6 new_pc) ]]).
+
+(* il classico push *)
+definition aux_push ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λval:byte8.
+ opt_map ?? (get_sp_reg m t s)
+  (* [SP] = val *)
+  (λSP_op.opt_map ?? (memory_filter_write m t s SP_op val)
+   (* SP -- *)
+   (λs_tmp1.opt_map ?? (set_sp_reg m t s_tmp1 (pred_w16 SP_op))
+    (λs_tmp2.Some ? s_tmp2))).
+
+(* il classico pop *)
+(* NB: l'incremento di SP deve essere filtrato dalla ALU, quindi get(set(SP)) *)
+definition aux_pop ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+ opt_map ?? (get_sp_reg m t s)
+  (* SP ++ *)
+  (λSP_op.opt_map ?? (set_sp_reg m t s (succ_w16 SP_op))
+   (λs_tmp1.opt_map ?? (get_sp_reg m t s_tmp1)
+    (* val = [SP] *)
+    (λSP_op'.opt_map ?? (memory_filter_read m t s_tmp1 SP_op')
+     (λval.Some ? (pairT ?? s_tmp1 val))))).
+
+(* CCR corrisponde a V11HINZC e cmq 1 se un flag non esiste *)
+(* i flag mantengono posizione costante nelle varie ALU, e se non sono
+   implementati corrispondono a 1 *)
+definition aux_get_CCR ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+let V_comp ≝ match get_v_flag m t s with
+ [ None ⇒ 〈x8,x0〉  | Some V_val ⇒ match V_val with [ true ⇒ 〈x8,x0〉  | false ⇒ 〈x0,x0〉 ]] in
+let H_comp ≝ match get_h_flag m t s with
+ [ None ⇒ 〈x1,x0〉  | Some H_val ⇒ match H_val with [ true ⇒ 〈x1,x0〉  | false ⇒ 〈x0,x0〉 ]] in
+let I_comp ≝ match get_i_flag m t s with
+ [ None ⇒ 〈x0,x8〉  | Some I_val ⇒ match I_val with [ true ⇒ 〈x0,x8〉  | false ⇒ 〈x0,x0〉 ]] in
+let N_comp ≝ match get_n_flag m t s with
+ [ None ⇒ 〈x0,x4〉  | Some N_val ⇒ match N_val with [ true ⇒ 〈x0,x4〉  | false ⇒ 〈x0,x0〉 ]] in
+let Z_comp ≝ match get_z_flag m t s with
+ [ true ⇒ 〈x0,x2〉  | false ⇒ 〈x0,x0〉 ] in
+let C_comp ≝ match get_c_flag m t s with
+ [ true ⇒ 〈x0,x1〉  | false ⇒ 〈x0,x0〉 ] in
+or_b8 〈x6,x0〉 (or_b8 V_comp (or_b8 H_comp (or_b8 I_comp (or_b8 N_comp (or_b8 Z_comp C_comp))))).
+
+(* CCR corrisponde a V11HINZC *)
+(* i flag mantengono posizione costante nelle varie ALU, e se non sono
+   implementati si puo' usare tranquillamente setweak *)
+definition aux_set_CCR ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λCCR:byte8.
+ let s_tmp1 ≝ set_c_flag m t s          (eq_b8 〈x0,x1〉 (and_b8 〈x0,x1〉 CCR)) in
+ let s_tmp2 ≝ set_z_flag m t s_tmp1     (eq_b8 〈x0,x2〉 (and_b8 〈x0,x2〉 CCR)) in
+ let s_tmp3 ≝ setweak_n_flag m t s_tmp2 (eq_b8 〈x0,x4〉 (and_b8 〈x0,x4〉 CCR)) in
+ let s_tmp4 ≝ setweak_i_flag m t s_tmp3 (eq_b8 〈x0,x8〉 (and_b8 〈x0,x8〉 CCR)) in
+ let s_tmp5 ≝ setweak_h_flag m t s_tmp4 (eq_b8 〈x1,x0〉 (and_b8 〈x1,x0〉 CCR)) in
+ let s_tmp6 ≝ setweak_v_flag m t s_tmp5 (eq_b8 〈x8,x0〉 (and_b8 〈x8,x0〉 CCR)) in
+ s_tmp6.
+
+(* **************** *)
+(* LOGICA DELLA ALU *)
+(* **************** *)
+
+(* A = A + M + C *)
+definition execute_ADC ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_ADC_ADD_aux m t s i cur_pc true (λA_op.λM_op.λC_op.plus_b8 A_op M_op C_op).
+
+(* A = A + M *)
+definition execute_ADD ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_ADC_ADD_aux m t s i cur_pc true (λA_op.λM_op.λC_op.plus_b8 A_op M_op false).
+
+(* SP += extended M *)
+definition execute_AIS ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (multi_mode_load true m t s cur_pc i)
+  (λS_M_PC.match S_M_PC with
+   [ tripleT s_tmp1 M_op new_pc ⇒
+   opt_map ?? (get_sp_reg m t s_tmp1)
+    (* SP += extended M *)
+    (λSP_op.opt_map ?? (set_sp_reg m t s_tmp1 (plus_w16nc SP_op (byte_extension M_op)))
+     (λs_tmp2.Some ? (pairT ?? s_tmp2 new_pc))) ]).
+
+(* H:X += extended M *)
+definition execute_AIX ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (multi_mode_load true m t s cur_pc i)
+  (λS_M_PC.match S_M_PC with
+   [ tripleT s_tmp1 M_op new_pc ⇒
+   opt_map ?? (get_indX_16_reg m t s_tmp1)
+    (* H:X += extended M *)
+    (λHX_op.opt_map ?? (set_indX_16_reg m t s_tmp1 (plus_w16nc HX_op (byte_extension M_op)))
+     (λs_tmp2.Some ? (pairT ?? s_tmp2 new_pc))) ]).
+
+(* A = A & M *)
+definition execute_AND ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_AND_BIT_EOR_ORA_aux m t s i cur_pc true and_b8.
+
+(* M = C' <- rcl M <- 0 *)
+definition execute_ASL ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_ASL_ASR_LSR_ROL_ROR_aux m t s i cur_pc (λM_op.λC_op.rcl_b8 M_op false).
+
+(* M = M7 -> rcr M -> C' *)
+definition execute_ASR ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_ASL_ASR_LSR_ROL_ROR_aux m t s i cur_pc (λM_op.λC_op.rcr_b8 M_op (MSB_b8 M_op)).
+
+(* if C=0, branch *) 
+definition execute_BCC ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_any_BRANCH m t s i cur_pc (⊖(get_c_flag m t s)).
+
+(* Mn = 0 *)
+definition execute_BCLRn ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_BCLRn_BSETn_aux m t s i cur_pc 〈x0,x0〉.
+
+(* if C=1, branch *) 
+definition execute_BCS ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_any_BRANCH m t s i cur_pc (get_c_flag m t s).
+
+(* if Z=1, branch *)
+definition execute_BEQ ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_any_BRANCH m t s i cur_pc (get_z_flag m t s).
+
+(* if N⊙V=0, branch *)
+definition execute_BGE ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_n_flag m t s)
+  (λN_op.opt_map ?? (get_v_flag m t s)
+   (λV_op.execute_any_BRANCH m t s i cur_pc (⊖(N_op ⊙ V_op)))).
+
+(* BGND mode *)
+definition execute_BGND ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ Some ? (pairT ?? s cur_pc).
+
+(* if Z|N⊙V=0, branch *)
+definition execute_BGT ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_n_flag m t s)
+  (λN_op.opt_map ?? (get_v_flag m t s)
+   (λV_op.execute_any_BRANCH m t s i cur_pc (⊖((get_z_flag m t s) ⊕ (N_op ⊙ V_op))))).
+
+(* if H=0, branch *)
+definition execute_BHCC ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_h_flag m t s)
+  (λH_op.execute_any_BRANCH m t s i cur_pc (⊖H_op)).
+
+(* if H=1, branch *)
+definition execute_BHCS ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_h_flag m t s)
+  (λH_op.execute_any_BRANCH m t s i cur_pc H_op).
+
+(* if C|Z=0, branch *)
+definition execute_BHI ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_any_BRANCH m t s i cur_pc (⊖((get_c_flag m t s) ⊕ (get_z_flag m t s))).
+
+(* if nIRQ=1, branch NB: irqflag e' un negato del pin *)
+definition execute_BIH ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_irq_flag m t s)
+  (λIRQ_op.execute_any_BRANCH m t s i cur_pc (⊖IRQ_op)).
+
+(* if nIRQ=0, branch NB: irqflag e' un negato del pin *)
+definition execute_BIL ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_irq_flag m t s)
+  (λIRQ_op.execute_any_BRANCH m t s i cur_pc IRQ_op).
+
+(* flags = A & M *)
+definition execute_BIT ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_AND_BIT_EOR_ORA_aux m t s i cur_pc false and_b8.
+
+(* if Z|N⊙V=1, branch *)
+definition execute_BLE ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_n_flag m t s)
+  (λN_op.opt_map ?? (get_v_flag m t s)
+   (λV_op.execute_any_BRANCH m t s i cur_pc ((get_z_flag m t s) ⊕ (N_op ⊙ V_op)))).
+
+(* if C|Z=1, branch *)
+definition execute_BLS ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_any_BRANCH m t s i cur_pc ((get_c_flag m t s) ⊕ (get_z_flag m t s)).
+
+(* if N⊙V=1, branch *)
+definition execute_BLT ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_n_flag m t s)
+  (λN_op.opt_map ?? (get_v_flag m t s)
+   (λV_op.execute_any_BRANCH m t s i cur_pc (N_op ⊙ V_op))).
+
+(* if I=0, branch *)
+definition execute_BMC ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_i_flag m t s)
+  (λI_op.execute_any_BRANCH m t s i cur_pc (⊖I_op)).
+
+(* if N=1, branch *)
+definition execute_BMI ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_n_flag m t s)
+  (λN_op.execute_any_BRANCH m t s i cur_pc N_op).
+
+(* if I=1, branch *)
+definition execute_BMS ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_i_flag m t s)
+  (λI_op.execute_any_BRANCH m t s i cur_pc I_op).
+
+(* if Z=0, branch *)
+definition execute_BNE ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_any_BRANCH m t s i cur_pc (⊖(get_z_flag m t s)).
+
+(* if N=0, branch *)
+definition execute_BPL ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_n_flag m t s)
+  (λN_op.execute_any_BRANCH m t s i cur_pc (⊖N_op)).
+
+(* branch always *)
+definition execute_BRA ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_any_BRANCH m t s i cur_pc true.
+
+(* if Mn=0 branch *)
+definition execute_BRCLRn ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_BRCLRn_BRSETn_aux m t s i cur_pc
+  (λMn_op.eq_b8 Mn_op 〈x0,x0〉).
+
+(* branch never... come se fosse un nop da 2 byte *)
+definition execute_BRN ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_any_BRANCH m t s i cur_pc false.
+
+(* if Mn=1 branch *)
+definition execute_BRSETn ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_BRCLRn_BRSETn_aux m t s i cur_pc
+  (λMn_op.⊖(eq_b8 Mn_op 〈x0,x0〉)).
+
+(* Mn = 1 *)
+definition execute_BSETn ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_BCLRn_BSETn_aux m t s i cur_pc 〈xF,xF〉.
+
+(* branch to subroutine *)
+(* HC05/HC08/HCS08 si appoggiano allo stack, RS08 a SPC *)
+definition execute_BSR ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t .λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (multi_mode_load true m t s cur_pc i)
+  (λS_M_PC.match S_M_PC with
+   [ tripleT s_tmp1 M_op new_pc ⇒ let aux ≝
+    (* push (new_pc low) *)
+    opt_map ?? (aux_push m t s_tmp1 (w16l new_pc))
+     (* push (new_pc high) *)
+     (λs_tmp2.opt_map ?? (aux_push m t s_tmp2 (w16h new_pc))
+      (* new_pc = new_pc + rel *)
+      (λs_tmp3.Some ? (pairT ?? s_tmp3 (branched_pc m t s_tmp3 new_pc M_op))))
+     in match m with
+    [ HC05 ⇒ aux | HC08 ⇒ aux | HCS08 ⇒ aux
+    | RS08 ⇒
+     (* SPC = new_pc *) 
+     opt_map ?? (set_spc_reg m t s_tmp1 new_pc)
+      (* new_pc = new_pc + rel *)
+      (λs_tmp2.Some ? (pairT ?? s_tmp2 (branched_pc m t s_tmp2 new_pc M_op)))
+    ]]).
+
+(* if A=M, branch *)
+definition execute_CBEQA ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (multi_mode_load false m t s cur_pc i)
+  (λS_M_PC.match S_M_PC with
+   [ tripleT s_tmp1 M_op new_pc ⇒
+    match M_op with
+     [ mk_word16 MH_op ML_op ⇒
+      (* if A=M, branch *)
+      match eq_b8 (get_acc_8_low_reg m t s_tmp1) MH_op with
+       (* new_pc = new_pc + rel *)
+       [ true ⇒ Some ? (pairT ?? s_tmp1 (branched_pc m t s_tmp1 new_pc ML_op))
+       (* new_pc = new_pc *)
+       | false ⇒ Some ? (pairT ?? s_tmp1 new_pc)
+       ]]]).
+
+(* if X=M, branch *)
+definition execute_CBEQX ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (multi_mode_load false m t s cur_pc i)
+  (λS_M_PC.match S_M_PC with
+   [ tripleT s_tmp1 M_op new_pc ⇒
+    match M_op with
+     [ mk_word16 MH_op ML_op ⇒
+      opt_map ?? (get_indX_8_low_reg m t s_tmp1)
+       (* if X=M, branch *)
+       (λX_op.match eq_b8 X_op MH_op with
+        (* new_pc = new_pc + rel *)
+        [ true ⇒ Some ? (pairT ?? s_tmp1 (branched_pc m t s_tmp1 new_pc ML_op))
+        (* new_pc = new_pc *)
+        | false ⇒ Some ? (pairT ?? s_tmp1 new_pc)
+        ])]]).
+
+(* C = 0 *)
+definition execute_CLC ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ Some ? (pairT ?? (set_c_flag m t s false) cur_pc).
+
+(* I = 0 *)
+definition execute_CLI ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (set_i_flag m t s false)
+  (λs_tmp.Some ? (pairT ?? s_tmp cur_pc)).
+
+(* M = 0 *)
+definition execute_CLR ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ (* M = 0 *)
+ opt_map ?? (multi_mode_write true m t s cur_pc i 〈x0,x0〉)
+  (λS_PC.match S_PC with
+   [ pairT s_tmp1 new_pc ⇒
+   (* Z = 1 *)
+   let s_tmp2 ≝ set_z_flag m t s_tmp1 true in
+   (* N = 0 *)
+   let s_tmp3 ≝ setweak_n_flag m t s_tmp2 false in
+   (* V = 0 *)
+   let s_tmp4 ≝ setweak_v_flag m t s_tmp3 false in
+   (* newpc = nextpc *)
+   Some ? (pairT ?? s_tmp4 new_pc) ]).
+
+(* flags = A - M *)
+definition execute_CMP ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_SBC_SUB_aux m t s i cur_pc false (λA_op.λM_op.λC_op.plus_b8 A_op (compl_b8 M_op) false). 
+
+(* M = not M *)
+definition execute_COM ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_COM_DEC_INC_NEG_aux m t s i cur_pc not_b8
+ (* fV = 0 *)
+ (λM7.λR7.false)
+ (* fC = 1 *)
+ (λC_op.λR_op.true).
+
+(* flags = H:X - M *)
+definition execute_CPHX ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (multi_mode_load false m t s cur_pc i)
+  (λS_M_PC.match S_M_PC with
+   [ tripleT s_tmp1 M_op new_pc ⇒
+    opt_map ?? (get_indX_16_reg m t s_tmp1)
+     (λX_op. 
+      match plus_w16 X_op (compl_w16 M_op) false with
+       [ pairT R_op carry ⇒
+        let X15 ≝ MSB_w16 X_op in let M15 ≝ MSB_w16 M_op in let R15 ≝ MSB_w16 R_op in
+        (* Z = nR15&nR14&nR13&nR12&nR11&nR10&nR9&nR8&nR7&nR6&nR5&nR4&nR3&nR2&nR1&nR0 *)
+        let s_tmp2 ≝ set_z_flag m t s_tmp1 (eq_w16 R_op 〈〈x0,x0〉:〈x0,x0〉〉) in
+        (* C = nX15&M15 | M15&R15 | R15&nX15 *)
+        let s_tmp3 ≝ set_c_flag m t s_tmp2 (((⊖X15)⊗M15) ⊕ (M15⊗R15) ⊕ (R15⊗(⊖X15))) in
+        (* N = R15 *) 
+        let s_tmp4 ≝ setweak_n_flag m t s_tmp3 R15 in
+        (* V = X15&nM15&nR15 | nX15&M15&R15 *)
+        let s_tmp5 ≝ setweak_v_flag m t s_tmp4 ((X15⊗(⊖M15)⊗(⊖R15)) ⊕ ((⊖X15)⊗M15⊗R15)) in
+        (* newpc = nextpc *)
+        Some ? (pairT ?? s_tmp5 new_pc) ] ) ]).
+
+(* flags = X - M *)
+definition execute_CPX ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (multi_mode_load true m t s cur_pc i)
+  (λS_M_PC.match S_M_PC with
+   [ tripleT s_tmp1 M_op new_pc ⇒
+    opt_map ?? (get_indX_8_low_reg m t s_tmp1)
+     (λX_op. 
+      match plus_b8 X_op (compl_b8 M_op) false with
+       [ pairT R_op carry ⇒
+        let X7 ≝ MSB_b8 X_op in let M7 ≝ MSB_b8 M_op in let R7 ≝ MSB_b8 R_op in
+        (* Z = nR7&nR6&nR5&nR4&nR3&nR2&nR1&nR0 *)
+        let s_tmp2 ≝ set_z_flag m t s_tmp1 (eq_b8 R_op 〈x0,x0〉) in
+        (* C = nX7&M7 | M7&R7 | R7&nX7 *)
+        let s_tmp3 ≝ set_c_flag m t s_tmp2 (((⊖X7)⊗M7) ⊕ (M7⊗R7) ⊕ (R7⊗(⊖X7))) in
+        (* N = R7 *) 
+        let s_tmp4 ≝ setweak_n_flag m t s_tmp3 R7 in
+        (* V = X7&nM7&nR7 | nX7&M7&R7 *)
+        let s_tmp5 ≝ setweak_v_flag m t s_tmp4 ((X7⊗(⊖M7)⊗(⊖R7)) ⊕ ((⊖X7)⊗M7⊗R7)) in
+        (* newpc = nextpc *)
+        Some ? (pairT ?? s_tmp5 new_pc) ] ) ]).
+
+(* decimal adjiust A *)
+(* per i dettagli vedere daa_b8 (modulo byte8) *)
+definition execute_DAA ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_h_flag m t s)
+  (λH.
+   let M_op ≝ get_acc_8_low_reg m t s in
+   match daa_b8 H (get_c_flag m t s) M_op with
+    [ pairT R_op carry ⇒
+     (* A = R *)
+     let s_tmp1 ≝ set_acc_8_low_reg m t s R_op in
+     (* Z = nR7&nR6&nR5&nR4&nR3&nR2&nR1&nR0 *)
+     let s_tmp2 ≝ set_z_flag m t s_tmp1 (eq_b8 R_op 〈x0,x0〉) in
+     (* C = carry *)
+     let s_tmp3 ≝ set_c_flag m t s_tmp2 carry in
+     (* N = R7 *) 
+     let s_tmp4 ≝ setweak_n_flag m t s_tmp3 (MSB_b8 R_op) in
+     (* V = M7 ⊙ R7 *)
+     let s_tmp5 ≝ setweak_v_flag m t s_tmp4 ((MSB_b8 M_op) ⊙ (MSB_b8 R_op)) in
+     (* newpc = curpc *)
+     Some ? (pairT ?? s_tmp5 cur_pc) ]).
+
+(* if (--M)≠0, branch *)
+definition execute_DBNZ ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (multi_mode_load false m t s cur_pc i)
+  (λS_M_PC.match S_M_PC with
+   [ tripleT s_tmp1 M_op new_pc ⇒
+    match M_op with
+     [ mk_word16 MH_op ML_op ⇒
+     (* --M *)
+     let MH_op' ≝ pred_b8 MH_op in
+     opt_map ?? (multi_mode_write true m t s_tmp1 cur_pc i MH_op')
+      (λS_PC.match S_PC with
+       [ pairT s_tmp2 _ ⇒
+        (* if (--M)≠0, branch *)
+        match eq_b8 MH_op' 〈x0,x0〉 with
+         (* new_pc = new_pc *)
+         [ true ⇒ Some ? (pairT ?? s_tmp2 new_pc)
+         (* new_pc = new_pc + rel *)
+         | false ⇒ Some ? (pairT ?? s_tmp2 (branched_pc m t s_tmp2 new_pc ML_op)) ]])]]).
+
+(* M = M - 1 *)
+definition execute_DEC ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_COM_DEC_INC_NEG_aux m t s i cur_pc pred_b8
+ (* fV = M7&nR7 *)
+ (λM7.λR7.M7⊗(⊖R7))
+ (* fC = C *)
+ (λC_op.λR_op.C_op).
+
+(* A = H:A/X, H = H:AmodX se non c'e' overflow, altrimenti invariati *)
+(* per i dettagli vedere div_b8 (modulo word16) *)
+definition execute_DIV ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_indX_8_high_reg m t s)
+  (λH_op.opt_map ?? (get_indX_8_low_reg m t s)
+   (λX_op.match div_b8 〈H_op:(get_acc_8_low_reg m t s)〉 X_op with
+    [ tripleT quoz rest overflow ⇒
+     (* C = overflow *)
+     let s_tmp1 ≝ set_c_flag m t s overflow in
+     (* A = A o H:A/X *)
+     let s_tmp2 ≝ match overflow with
+      [ true ⇒ s_tmp1
+      | false ⇒ set_acc_8_low_reg m t s_tmp1 quoz ] in
+     (* Z = nA7&nA6&nA5&nA4&nA3&nA2&nA1&nA0 *)
+     (* NB: che A sia cambiato o no, lo testa *)
+     let s_tmp3 ≝ set_z_flag m t s_tmp2 (eq_b8 (get_acc_8_low_reg m t s_tmp2) 〈x0,x0〉) in
+     (* H = H o H:AmodX *)
+     opt_map ?? (match overflow with
+                 [ true ⇒ Some ? s_tmp3
+                 | false ⇒ set_indX_8_high_reg m t s_tmp3 rest])
+      (λs_tmp4.Some ? (pairT ?? s_tmp4 cur_pc)) ])).
+
+(* A = A ⊙ M *)
+definition execute_EOR ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_AND_BIT_EOR_ORA_aux m t s i cur_pc true xor_b8.
+
+(* M = M + 1 *)
+definition execute_INC ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_COM_DEC_INC_NEG_aux m t s i cur_pc succ_b8
+ (* fV = nM7&R7 *)
+ (λM7.λR7.(⊖M7)⊗R7)
+ (* fC = C *)
+ (λC_op.λR_op.C_op).
+
+(* jmp, il nuovo indirizzo e' una WORD *)
+definition execute_JMP ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (multi_mode_load false m t s cur_pc i)
+  (λS_M_PC.
+   (* newpc = M_op *)
+   Some ? (pairT ?? (fst3T ??? S_M_PC) (snd3T ??? S_M_PC))).
+
+(* jump to subroutine *)
+(* HC05/HC08/HCS08 si appoggiano allo stack, RS08 a SPC *)
+definition execute_JSR ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (multi_mode_load false m t s cur_pc i)
+  (λS_M_PC.match S_M_PC with
+   [ tripleT s_tmp1 M_op new_pc ⇒ let aux ≝
+    (* push (new_pc low) *)
+    opt_map ?? (aux_push m t s_tmp1 (w16l new_pc))
+     (* push (new_pc high) *)
+     (λs_tmp2.opt_map ?? (aux_push m t s_tmp2 (w16h new_pc))
+      (* newpc = M_op *)
+      (λs_tmp3.Some ? (pairT ?? s_tmp3 M_op)))
+     in match m with
+    [ HC05 ⇒ aux | HC08 ⇒ aux | HCS08 ⇒ aux
+    | RS08 ⇒
+     (* SPC = new_pc *) 
+     opt_map ?? (set_spc_reg m t s_tmp1 new_pc)
+      (* newpc = M_op *)
+      (λs_tmp2.Some ? (pairT ?? s_tmp2 M_op))
+    ]]).
+
+(* A = M *)
+definition execute_LDA ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (multi_mode_load true m t s cur_pc i)
+  (λS_M_PC.match S_M_PC with
+   [ tripleT s_tmp1 M_op new_pc ⇒
+    (* A = M *) 
+    let s_tmp2 ≝ set_acc_8_low_reg m t s_tmp1 M_op in
+    (* Z = nR7&nR6&nR5&nR4&nR3&nR2&nR1&nR0 *)
+    let s_tmp3 ≝ set_z_flag m t s_tmp2 (eq_b8 M_op 〈x0,x0〉) in
+    (* N = R7 *) 
+    let s_tmp4 ≝ setweak_n_flag m t s_tmp3 (MSB_b8 M_op) in
+    (* V = 0 *) 
+    let s_tmp5 ≝ setweak_v_flag m t s_tmp4 false in
+    (* newpc = nextpc *)
+    Some ? (pairT ?? s_tmp5 new_pc) ]).
+
+(* H:X = M *)
+definition execute_LDHX ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (multi_mode_load false m t s cur_pc i)
+  (λS_M_PC.match S_M_PC with
+   [ tripleT s_tmp1 M_op new_pc ⇒
+    opt_map ?? (set_indX_16_reg m t s_tmp1 M_op)
+     (λs_tmp2.
+      (* Z = nR15&nR14&nR13nR12&nR11&nR10&nR9&nR8nR7&nR6&nR5&nR4&nR3&nR2&nR1&nR0 *)
+      let s_tmp3 ≝ set_z_flag m t s_tmp2 (eq_w16 M_op 〈〈x0,x0〉:〈x0,x0〉〉) in
+      (* N = R15 *)
+      let s_tmp4 ≝ setweak_n_flag m t s_tmp3 (MSB_w16 M_op) in
+      (* V = 0 *) 
+      let s_tmp5 ≝ setweak_v_flag m t s_tmp4 false in
+      (* newpc = nextpc *)
+      Some ? (pairT ?? s_tmp5 new_pc)) ]).
+
+(* X = M *)
+definition execute_LDX ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (multi_mode_load true m t s cur_pc i)
+  (λS_M_PC.match S_M_PC with
+   [ tripleT s_tmp1 M_op new_pc ⇒
+    opt_map ?? (set_indX_8_low_reg m t s_tmp1 M_op)
+     (λs_tmp2.
+      (* Z = nR7&nR6&nR5&nR4&nR3&nR2&nR1&nR0 *)
+      let s_tmp3 ≝ set_z_flag m t s_tmp2 (eq_b8 M_op 〈x0,x0〉) in
+      (* N = R7 *)
+      let s_tmp4 ≝ setweak_n_flag m t s_tmp3 (MSB_b8 M_op) in
+      (* V = 0 *) 
+      let s_tmp5 ≝ setweak_v_flag m t s_tmp4 false in
+      (* newpc = nextpc *)
+      Some ? (pairT ?? s_tmp5 new_pc)) ]).
+
+(* M = 0 -> rcr M -> C' *)
+definition execute_LSR ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_ASL_ASR_LSR_ROL_ROR_aux m t s i cur_pc (λM_op.λC_op.rcr_b8 M_op false).
+
+(* M2 = M1 *)
+definition execute_MOV ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ (* R_op = M1 *)
+ opt_map ?? (multi_mode_load true m t s cur_pc i)
+  (λS_R_PC.match S_R_PC with
+   [ tripleT s_tmp1 R_op tmp_pc ⇒
+    (* M2 = R_op *)
+    opt_map ?? (multi_mode_write true m t s_tmp1 tmp_pc i R_op)
+     (λS_PC.match S_PC with
+      [ pairT s_tmp2 new_pc ⇒
+       (* Z = nR7&nR6&nR5&nR4&nR3&nR2&nR1&nR0 *)
+       let s_tmp3 ≝ set_z_flag m t s_tmp2 (eq_b8 R_op 〈x0,x0〉) in
+       (* N = R7 *)
+       let s_tmp4 ≝ setweak_n_flag m t s_tmp3 (MSB_b8 R_op) in
+       (* V = 0 *) 
+       let s_tmp5 ≝ setweak_v_flag m t s_tmp4 false in
+       (* newpc = nextpc *)
+       Some ? (pairT ?? s_tmp5 new_pc)])]).
+
+(* X:A = X * A *)
+definition execute_MUL ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_indX_8_low_reg m t s)
+  (λX_op.let R_op ≝ mul_b8 X_op (get_acc_8_low_reg m t s) in
+   opt_map ?? (set_indX_8_low_reg m t s (w16h R_op))
+    (λs_tmp.Some ? (pairT ?? (set_acc_8_low_reg m t s_tmp (w16l R_op)) cur_pc))).
+
+(* M = compl M *)
+definition execute_NEG ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_COM_DEC_INC_NEG_aux m t s i cur_pc compl_b8
+ (* fV = M7&R7 *)
+ (λM7.λR7.M7⊗R7)
+ (* fC = R7|R6|R5|R4|R3|R2|R1|R0 *)
+ (λC_op.λR_op.⊖(eq_b8 R_op 〈x0,x0〉)).
+
+(* nulla *)
+definition execute_NOP ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ Some ? (pairT ?? s cur_pc).
+
+(* A = (mk_byte8 (b8l A) (b8h A)) *)
+(* cioe' swap del nibble alto/nibble basso di A *)
+definition execute_NSA ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ match get_acc_8_low_reg m t s with [ mk_byte8 ah al ⇒
+  (* A = (mk_byte8 (b8l A) (b8h A)) *)
+  Some ? (pairT ?? (set_acc_8_low_reg m t s 〈al,ah〉) cur_pc) ].
+
+(* A = A | M *)
+definition execute_ORA ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_AND_BIT_EOR_ORA_aux m t s i cur_pc true or_b8.
+
+(* push A *)
+definition execute_PSHA ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (aux_push m t s (get_acc_8_low_reg m t s))
+  (λs_tmp1.Some ? (pairT ?? s_tmp1 cur_pc)).
+
+(* push H *)
+definition execute_PSHH ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_indX_8_high_reg m t s)
+  (λH_op.opt_map ?? (aux_push m t s H_op)
+   (λs_tmp1.Some ? (pairT ?? s_tmp1 cur_pc))).
+
+(* push X *)
+definition execute_PSHX ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_indX_8_low_reg m t s)
+  (λH_op.opt_map ?? (aux_push m t s H_op)
+   (λs_tmp1.Some ? (pairT ?? s_tmp1 cur_pc))).
+
+(* pop A *)
+definition execute_PULA ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (aux_pop m t s)
+  (λS_and_A.match S_and_A with [ pairT s_tmp1 A_op ⇒
+   Some ? (pairT ?? (set_acc_8_low_reg m t s_tmp1 A_op) cur_pc) ]).
+
+(* pop H *)
+definition execute_PULH ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (aux_pop m t s)
+  (λS_and_H.match S_and_H with [ pairT s_tmp1 H_op ⇒
+   opt_map ?? (set_indX_8_high_reg m t s_tmp1 H_op)
+    (λs_tmp2.Some ? (pairT ?? s_tmp2 cur_pc))]).
+
+(* pop X *)
+definition execute_PULX ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (aux_pop m t s)
+  (λS_and_X.match S_and_X with [ pairT s_tmp1 X_op ⇒
+   opt_map ?? (set_indX_8_low_reg m t s_tmp1 X_op)
+    (λs_tmp2.Some ? (pairT ?? s_tmp2 cur_pc))]).
+
+(* M = C' <- rcl M <- C *)
+definition execute_ROL ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_ASL_ASR_LSR_ROL_ROR_aux m t s i cur_pc (λM_op.λC_op.rcl_b8 M_op C_op).
+
+(* M = C -> rcr M -> C' *)
+definition execute_ROR ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_ASL_ASR_LSR_ROL_ROR_aux m t s i cur_pc (λM_op.λC_op.rcr_b8 M_op C_op).
+
+(* SP = 0xuuFF *)
+(* lascia inalterato il byte superiore di SP *)
+definition execute_RSP ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_sp_reg m t s)
+  (λSP_op.match SP_op with [ mk_word16 sph spl ⇒
+   opt_map ?? (set_sp_reg m t s 〈sph:〈xF,xF〉〉)
+    (λs_tmp.Some ? (pairT ?? s_tmp cur_pc))]).
+
+(* return from interrupt *)
+definition execute_RTI ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ (* pop (CCR) *)
+ opt_map ?? (aux_pop m t s)
+  (λS_and_CCR.match S_and_CCR with [ pairT s_tmp1 CCR_op ⇒
+   let s_tmp2 ≝ aux_set_CCR m t s_tmp1 CCR_op in
+   (* pop (A) *)
+   opt_map ?? (aux_pop m t s_tmp2)
+    (λS_and_A.match S_and_A with [ pairT s_tmp3 A_op ⇒
+     let s_tmp4 ≝ set_acc_8_low_reg m t s_tmp3 A_op in
+     (* pop (X) *)
+     opt_map ?? (aux_pop m t s_tmp4)
+      (λS_and_X.match S_and_X with [ pairT s_tmp5 X_op ⇒
+       opt_map ?? (set_indX_8_low_reg m t s_tmp5 X_op)
+        (* pop (PC high) *)
+        (λs_tmp6.opt_map ?? (aux_pop m t s_tmp6)
+         (λS_and_PCH.match S_and_PCH with [ pairT s_tmp7 PCH_op ⇒
+          (* pop (PC low) *)
+          opt_map ?? (aux_pop m t s_tmp7)
+           (λS_and_PCL.match S_and_PCL with [ pairT s_tmp8 PCL_op ⇒
+            Some ? (pairT ?? s_tmp8 〈PCH_op:PCL_op〉)])]))])])]).
+
+(* return from subroutine *)
+(* HC05/HC08/HCS08 si appoggia allo stack, RS08 si appoggia a SPC *)
+definition execute_RTS ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ let aux ≝
+  (* pop (PC high) *)
+  opt_map ?? (aux_pop m t s)
+   (λS_and_PCH.match S_and_PCH with [ pairT s_tmp1 PCH_op ⇒
+    (* pop (PC low) *)
+    opt_map ?? (aux_pop m t s_tmp1)
+     (λS_and_PCL.match S_and_PCL with [ pairT s_tmp2 PCL_op ⇒
+      Some ? (pairT ?? s_tmp2 〈PCH_op:PCL_op〉)])])
+ in match m with
+  [ HC05 ⇒ aux | HC08 ⇒ aux | HCS08 ⇒ aux
+  | RS08 ⇒
+   (* new_pc = SPC *)
+   opt_map ?? (get_spc_reg m t s)
+    (λSPC_op.Some ? (pairT ?? s SPC_op))
+  ].
+
+(* A = A - M - C *)
+definition execute_SBC ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_SBC_SUB_aux m t s i cur_pc true
+ (λA_op.λM_op.λC_op.match plus_b8 A_op (compl_b8 M_op) false with
+  [ pairT resb resc ⇒ match C_op with
+   [ true ⇒ plus_b8 resb 〈xF,xF〉 resc
+   | false ⇒ pairT ?? resb resc ]]).
+
+(* C = 1 *)
+definition execute_SEC ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ Some ? (pairT ?? (set_c_flag m t s true) cur_pc).
+
+(* I = 1 *)
+definition execute_SEI ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (set_i_flag m t s true)
+  (λs_tmp.Some ? (pairT ?? s_tmp cur_pc)).
+
+(* swap SPCh,A *)
+(* senso: nell'RS08 SPC non e' accessibile direttamente e come si possono
+          fare subroutine annidate se RA (return address) e' salvato sempre in SPC?
+          occore accedere a SPC e salvarne il contenuto *)
+definition execute_SHA ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_spc_reg m t s)
+  (λSPC_op.opt_map ?? (set_spc_reg m t s 〈(get_acc_8_low_reg m t s):(w16l SPC_op)〉)
+   (λs_tmp1.Some ? (pairT ?? (set_acc_8_low_reg m t s_tmp1 (w16h SPC_op)) cur_pc))).
+
+(* swap SPCl,A *)
+(* senso: nell'RS08 SPC non e' accessibile direttamente e come si possono
+          fare subroutine annidate se RA (return address) e' salvato sempre in SPC?
+          occore accedere a SPC e salvarne il contenuto *)
+definition execute_SLA ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_spc_reg m t s)
+  (λSPC_op.opt_map ?? (set_spc_reg m t s 〈(w16h SPC_op):(get_acc_8_low_reg m t s)〉)
+   (λs_tmp1.Some ? (pairT ?? (set_acc_8_low_reg m t s_tmp1 (w16l SPC_op)) cur_pc))).
+
+(* M = A *)
+definition execute_STA ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ (* M = A *)
+ let A_op ≝ (get_acc_8_low_reg m t s) in
+ opt_map ?? (multi_mode_write true m t s cur_pc i A_op)
+  (λS_op_and_PC.match S_op_and_PC with
+   [ pairT s_tmp1 new_pc ⇒
+   (* Z = nA7&nA6&nA5&nA4&nA3&nA2&nA1&nA0 *)
+   let s_tmp2 ≝ set_z_flag m t s_tmp1 (eq_b8 A_op 〈x0,x0〉) in
+   (* N = A7 *)
+   let s_tmp3 ≝ setweak_n_flag m t s_tmp2 (MSB_b8 A_op) in
+   (* V = 0 *)
+   let s_tmp4 ≝ setweak_v_flag m t s_tmp3 false in
+   (* newpc = nextpc *)
+   Some ? (pairT ?? s_tmp4 new_pc) ]).
+
+(* M = H:X *)
+definition execute_STHX ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ (* M = H:X *)
+ opt_map ?? (get_indX_16_reg m t s)
+  (λX_op.opt_map ?? (multi_mode_write false m t s cur_pc i X_op)
+   (λS_op_and_PC.match S_op_and_PC with
+    [ pairT s_tmp1 new_pc ⇒
+     (* Z = nR15&nR14&nR13nR12&nR11&nR10&nR9&nR8nR7&nR6&nR5&nR4&nR3&nR2&nR1&nR0 *)
+     let s_tmp2 ≝ set_z_flag m t s_tmp1 (eq_w16 X_op 〈〈x0,x0〉:〈x0,x0〉〉) in
+     (* N = R15 *)
+     let s_tmp3 ≝ setweak_n_flag m t s_tmp2 (MSB_w16 X_op) in
+     (* V = 0 *)
+     let s_tmp4 ≝ setweak_v_flag m t s_tmp3 false in
+     (* newpc = nextpc *)
+      Some ? (pairT ?? s_tmp4 new_pc) ])).
+
+(* I = 0 *)
+definition execute_STOP ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ Some ? (pairT ?? (setweak_i_flag m t s false) cur_pc).
+
+(* M = X *)
+definition execute_STX ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ (* M = X *)
+ opt_map ?? (get_indX_8_low_reg m t s)
+  (λX_op.opt_map ?? (multi_mode_write true m t s cur_pc i X_op)
+   (λS_op_and_PC.match S_op_and_PC with
+    [ pairT s_tmp1 new_pc ⇒
+     (* Z = nR7&nR6&nR5&nR4&nR3&nR2&nR1&nR0 *)
+     let s_tmp2 ≝ set_z_flag m t s_tmp1 (eq_b8 X_op 〈x0,x0〉) in
+     (* N = R7 *)
+     let s_tmp3 ≝ setweak_n_flag m t s_tmp2 (MSB_b8 X_op) in
+     (* V = 0 *)
+     let s_tmp4 ≝ setweak_v_flag m t s_tmp3 false in
+     (* newpc = nextpc *)
+     Some ? (pairT ?? s_tmp4 new_pc) ])).
+
+(* A = A - M *)
+definition execute_SUB ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ execute_SBC_SUB_aux m t s i cur_pc true (λA_op.λM_op.λC_op.plus_b8 A_op (compl_b8 M_op) false). 
+
+(* software interrupt *)
+definition execute_SWI ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ (* indirizzo da cui caricare il nuovo pc *)
+ let vector ≝ get_pc_reg m t (set_pc_reg m t s 〈〈xF,xF〉:〈xF,xC〉〉) in
+ (* push (cur_pc low) *)
+ opt_map ?? (aux_push m t s (w16l cur_pc))
+  (* push (cur_pc high *)
+  (λs_tmp1.opt_map ?? (aux_push m t s_tmp1 (w16h cur_pc))
+   (λs_tmp2.opt_map ?? (get_indX_8_low_reg m t s_tmp2)
+    (* push (X) *)
+    (λX_op.opt_map ?? (aux_push m t s_tmp2 X_op)
+     (* push (A) *)
+     (λs_tmp3.opt_map ?? (aux_push m t s_tmp3 (get_acc_8_low_reg m t s_tmp3))
+      (* push (CCR) *)
+      (λs_tmp4.opt_map ?? (aux_push m t s_tmp4 (aux_get_CCR m t s_tmp4))
+       (* I = 1 *)
+       (λs_tmp5.opt_map ?? (set_i_flag m t s_tmp5 true)
+        (* load from vector high *)
+        (λs_tmp6.opt_map ?? (memory_filter_read m t s_tmp6 vector)
+         (* load from vector low *)
+         (λaddrh.opt_map ?? (memory_filter_read m t s_tmp6 (succ_w16 vector))
+          (* newpc = [vector] *)
+          (λaddrl.Some ? (pairT ?? s_tmp6 〈addrh:addrl〉)))))))))).
+
+(* flags = A *)
+definition execute_TAP ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ Some ? (pairT ?? (aux_set_CCR m t s (get_acc_8_low_reg m t s)) cur_pc). 
+
+(* X = A *)
+definition execute_TAX ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (set_indX_8_low_reg m t s (get_acc_8_low_reg m t s))
+  (λs_tmp.Some ? (pairT ?? s_tmp cur_pc)).
+
+(* A = flags *)
+definition execute_TPA ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ Some ? (pairT ?? (set_acc_8_low_reg m t s (aux_get_CCR m t s)) cur_pc).
+
+(* flags = M - 0 *)
+(* implementata senza richiamare la sottrazione, la modifica dei flag
+   e' immediata *)
+definition execute_TST ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (multi_mode_load true m t s cur_pc i)
+  (λS_M_PC.match S_M_PC with
+   [ tripleT s_tmp1 M_op new_pc ⇒
+    (* Z = nR7&nR6&nR5&nR4&nR3&nR2&nR1&nR0 *)
+    let s_tmp2 ≝ set_z_flag m t s_tmp1 (eq_b8 M_op 〈x0,x0〉) in
+    (* N = R7 *) 
+    let s_tmp3 ≝ setweak_n_flag m t s_tmp2 (MSB_b8 M_op) in
+    (* V = 0 *) 
+    let s_tmp4 ≝ setweak_v_flag m t s_tmp3 false in
+    (* newpc = nextpc *)
+    Some ? (pairT ?? s_tmp4 new_pc) ]).
+
+(* H:X = SP + 1 *)
+definition execute_TSX ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_sp_reg m t s )
+  (λSP_op.opt_map ?? (set_indX_16_reg m t s (succ_w16 SP_op))
+   (* H:X = SP + 1 *)
+   (λs_tmp.Some ? (pairT ?? s_tmp cur_pc))).
+
+(* A = X *)
+definition execute_TXA ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_indX_8_low_reg m t s)
+  (λX_op.Some ? (pairT ?? (set_acc_8_low_reg m t s X_op) cur_pc)).
+
+(* SP = H:X - 1 *)
+definition execute_TXS ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ opt_map ?? (get_indX_16_reg m t s )
+  (λX_op.opt_map ?? (set_sp_reg m t s (pred_w16 X_op))
+   (* SP = H:X - 1 *)
+   (λs_tmp.Some ? (pairT ?? s_tmp cur_pc))).
+
+(* I = 0 *)
+definition execute_WAIT ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λi:instr_mode.λcur_pc:word16.
+ Some ? (pairT ?? (setweak_i_flag m t s false) cur_pc).
+
+(* **** *)
+(* TICK *)
+(* **** *)
+
+(* enumerazione delle possibili modalita' di sospensione *)
+inductive susp_type : Type ≝
+  BGND_MODE: susp_type
+| STOP_MODE: susp_type
+| WAIT_MODE: susp_type.
+
+(* un tipo opzione ad hoc
+   - errore: errore+stato (seguira' reset o ??, cmq lo stato non va buttato)
+   - sospensione: sospensione+stato (seguira' resume o ??)
+   - ok: stato
+*)
+inductive tick_result (A:Type) : Type ≝
+  TickERR   : A → error_type → tick_result A
+| TickSUSP  : A → susp_type → tick_result A
+| TickOK    : A → tick_result A.
+
+(* sostanazialmente simula
+   - fetch/decode/execute
+   - l'esecuzione e' considerata atomica quindi nel caso di un'instruzione
+     da 3 cicli la successione sara'
+       ([fetch/decode] s,clk:None) →
+       (               s,clk:Some 1,pseudo,mode,3,cur_pc) →
+       (               s,clk:Some 2,pseudo,mode,3,cur_pc) →
+       ([execute]      s',clk:None) *)
+
+definition tick_execute ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+λpseudo:any_opcode m.λmode:instr_mode.λcur_pc:word16.
+ let abs_pseudo ≝ match pseudo with [ anyOP pseudo' ⇒ pseudo' ] in
+ let a_status_and_fetch ≝ match abs_pseudo with
+  [ ADC    ⇒ execute_ADC    m t s mode cur_pc (* add with carry *)
+  | ADD    ⇒ execute_ADD    m t s mode cur_pc (* add *)
+  | AIS    ⇒ execute_AIS    m t s mode cur_pc (* add immediate to SP *)
+  | AIX    ⇒ execute_AIX    m t s mode cur_pc (* add immediate to X *)
+  | AND    ⇒ execute_AND    m t s mode cur_pc (* and *)
+  | ASL    ⇒ execute_ASL    m t s mode cur_pc (* aritmetic shift left *)
+  | ASR    ⇒ execute_ASR    m t s mode cur_pc (* aritmetic shift right *)
+  | BCC    ⇒ execute_BCC    m t s mode cur_pc (* branch if C=0 *)
+  | BCLRn  ⇒ execute_BCLRn  m t s mode cur_pc (* clear bit n *)
+  | BCS    ⇒ execute_BCS    m t s mode cur_pc (* branch if C=1 *)
+  | BEQ    ⇒ execute_BEQ    m t s mode cur_pc (* branch if Z=1 *)
+  | BGE    ⇒ execute_BGE    m t s mode cur_pc (* branch if N⊙V=0 (great or equal) *)
+  | BGND   ⇒ execute_BGND   m t s mode cur_pc (* !!background mode!!*)
+  | BGT    ⇒ execute_BGT    m t s mode cur_pc (* branch if Z|N⊙V=0 clear (great) *)
+  | BHCC   ⇒ execute_BHCC   m t s mode cur_pc (* branch if H=0 *)
+  | BHCS   ⇒ execute_BHCS   m t s mode cur_pc (* branch if H=1 *)
+  | BHI    ⇒ execute_BHI    m t s mode cur_pc (* branch if C|Z=0, (higher) *)
+  | BIH    ⇒ execute_BIH    m t s mode cur_pc (* branch if nIRQ=1 *)
+  | BIL    ⇒ execute_BIL    m t s mode cur_pc (* branch if nIRQ=0 *)
+  | BIT    ⇒ execute_BIT    m t s mode cur_pc (* flag = and (bit test) *)
+  | BLE    ⇒ execute_BLE    m t s mode cur_pc (* branch if Z|N⊙V=1 (less or equal) *)
+  | BLS    ⇒ execute_BLS    m t s mode cur_pc (* branch if C|Z=1 (lower or same) *)
+  | BLT    ⇒ execute_BLT    m t s mode cur_pc (* branch if N⊙1=1 (less) *)
+  | BMC    ⇒ execute_BMC    m t s mode cur_pc (* branch if I=0 (interrupt mask clear) *)
+  | BMI    ⇒ execute_BMI    m t s mode cur_pc (* branch if N=1 (minus) *)
+  | BMS    ⇒ execute_BMS    m t s mode cur_pc (* branch if I=1 (interrupt mask set) *)
+  | BNE    ⇒ execute_BNE    m t s mode cur_pc (* branch if Z=0 *)
+  | BPL    ⇒ execute_BPL    m t s mode cur_pc (* branch if N=0 (plus) *)
+  | BRA    ⇒ execute_BRA    m t s mode cur_pc (* branch always *)
+  | BRCLRn ⇒ execute_BRCLRn m t s mode cur_pc (* branch if bit n clear *)
+  | BRN    ⇒ execute_BRN    m t s mode cur_pc (* branch never (nop) *)
+  | BRSETn ⇒ execute_BRSETn m t s mode cur_pc (* branch if bit n set *)
+  | BSETn  ⇒ execute_BSETn  m t s mode cur_pc (* set bit n *)
+  | BSR    ⇒ execute_BSR    m t s mode cur_pc (* branch to subroutine *)
+  | CBEQA  ⇒ execute_CBEQA  m t s mode cur_pc (* compare (A) and BEQ *)
+  | CBEQX  ⇒ execute_CBEQX  m t s mode cur_pc (* compare (X) and BEQ *)
+  | CLC    ⇒ execute_CLC    m t s mode cur_pc (* C=0 *)
+  | CLI    ⇒ execute_CLI    m t s mode cur_pc (* I=0 *)
+  | CLR    ⇒ execute_CLR    m t s mode cur_pc (* operand=0 *)
+  | CMP    ⇒ execute_CMP    m t s mode cur_pc (* flag = sub (compare A) *)
+  | COM    ⇒ execute_COM    m t s mode cur_pc (* not (1 complement) *)
+  | CPHX   ⇒ execute_CPHX   m t s mode cur_pc (* flag = sub (compare H:X) *)
+  | CPX    ⇒ execute_CPX    m t s mode cur_pc (* flag = sub (compare X) *)
+  | DAA    ⇒ execute_DAA    m t s mode cur_pc (* decimal adjust A *)
+  | DBNZ   ⇒ execute_DBNZ   m t s mode cur_pc (* dec and BNE *)
+  | DEC    ⇒ execute_DEC    m t s mode cur_pc (* operand=operand-1 (decrement) *)
+  | DIV    ⇒ execute_DIV    m t s mode cur_pc (* div *)
+  | EOR    ⇒ execute_EOR    m t s mode cur_pc (* xor *)
+  | INC    ⇒ execute_INC    m t s mode cur_pc (* operand=operand+1 (increment) *)
+  | JMP    ⇒ execute_JMP    m t s mode cur_pc (* jmp word [operand] *)
+  | JSR    ⇒ execute_JSR    m t s mode cur_pc (* jmp to subroutine *)
+  | LDA    ⇒ execute_LDA    m t s mode cur_pc (* load in A *)
+  | LDHX   ⇒ execute_LDHX   m t s mode cur_pc (* load in H:X *)
+  | LDX    ⇒ execute_LDX    m t s mode cur_pc (* load in X *)
+  | LSR    ⇒ execute_LSR    m t s mode cur_pc (* logical shift right *)
+  | MOV    ⇒ execute_MOV    m t s mode cur_pc (* move *)
+  | MUL    ⇒ execute_MUL    m t s mode cur_pc (* mul *)
+  | NEG    ⇒ execute_NEG    m t s mode cur_pc (* neg (2 complement) *)
+  | NOP    ⇒ execute_NOP    m t s mode cur_pc (* nop *)
+  | NSA    ⇒ execute_NSA    m t s mode cur_pc (* nibble swap A (al:ah <- ah:al) *)
+  | ORA    ⇒ execute_ORA    m t s mode cur_pc (* or *)
+  | PSHA   ⇒ execute_PSHA   m t s mode cur_pc (* push A *)
+  | PSHH   ⇒ execute_PSHH   m t s mode cur_pc (* push H *)
+  | PSHX   ⇒ execute_PSHX   m t s mode cur_pc (* push X *)
+  | PULA   ⇒ execute_PULA   m t s mode cur_pc (* pop A *)
+  | PULH   ⇒ execute_PULH   m t s mode cur_pc (* pop H *)
+  | PULX   ⇒ execute_PULX   m t s mode cur_pc (* pop X *)
+  | ROL    ⇒ execute_ROL    m t s mode cur_pc (* rotate left *)
+  | ROR    ⇒ execute_ROR    m t s mode cur_pc (* rotate right *)
+  | RSP    ⇒ execute_RSP    m t s mode cur_pc (* reset SP (0x00FF) *)
+  | RTI    ⇒ execute_RTI    m t s mode cur_pc (* return from interrupt *)
+  | RTS    ⇒ execute_RTS    m t s mode cur_pc (* return from subroutine *)
+  | SBC    ⇒ execute_SBC    m t s mode cur_pc (* sub with carry*)
+  | SEC    ⇒ execute_SEC    m t s mode cur_pc (* C=1 *)
+  | SEI    ⇒ execute_SEI    m t s mode cur_pc (* I=1 *)
+  | SHA    ⇒ execute_SHA    m t s mode cur_pc (* swap spc_high,A *)
+  | SLA    ⇒ execute_SLA    m t s mode cur_pc (* swap spc_low,A *)
+  | STA    ⇒ execute_STA    m t s mode cur_pc (* store from A *)
+  | STHX   ⇒ execute_STHX   m t s mode cur_pc (* store from H:X *)
+  | STOP   ⇒ execute_STOP   m t s mode cur_pc (* !!stop mode!! *)
+  | STX    ⇒ execute_STX    m t s mode cur_pc (* store from X *)
+  | SUB    ⇒ execute_SUB    m t s mode cur_pc (* sub *)
+  | SWI    ⇒ execute_SWI    m t s mode cur_pc (* software interrupt *)
+  | TAP    ⇒ execute_TAP    m t s mode cur_pc (* flag=A (transfer A to process status byte *)
+  | TAX    ⇒ execute_TAX    m t s mode cur_pc (* X=A (transfer A to X) *)
+  | TPA    ⇒ execute_TPA    m t s mode cur_pc (* A=flag (transfer process status byte to A) *)
+  | TST    ⇒ execute_TST    m t s mode cur_pc (* flag = sub (test) *)
+  | TSX    ⇒ execute_TSX    m t s mode cur_pc (* X:H=SP (transfer SP to H:X) *)
+  | TXA    ⇒ execute_TXA    m t s mode cur_pc (* A=X (transfer X to A) *)
+  | TXS    ⇒ execute_TXS    m t s mode cur_pc (* SP=X:H (transfer H:X to SP) *)
+  | WAIT   ⇒ execute_WAIT   m t s mode cur_pc (* !!wait mode!!*)
+  ] in match a_status_and_fetch with
+(* errore nell'execute (=caricamento argomenti)? riportato in output *)
+(* nessun avanzamento e clk a None *)
+   [ None ⇒ TickERR ? (set_clk_desc m t s (None ?)) ILL_EX_AD
+   | Some status_and_newpc ⇒
+(* aggiornamento centralizzato di pc e clk *)
+     match status_and_newpc with
+      [ pairT s_tmp1 new_pc ⇒
+       let s_tmp2 ≝ set_pc_reg m t s_tmp1 new_pc in
+       let s_tmp3 ≝ set_clk_desc m t s_tmp2 (None ?) in
+       let abs_magic ≝ magic_of_opcode abs_pseudo in
+(* distinzione fra le 4 modalita' possibili, normale/BGND/STOP/WAIT *)
+        match eq_b8 abs_magic (magic_of_opcode BGND) with
+         [ true ⇒ TickSUSP ? s_tmp3 BGND_MODE
+         | false ⇒ match eq_b8 abs_magic (magic_of_opcode STOP) with
+          [ true ⇒ TickSUSP ? s_tmp3 STOP_MODE
+          | false ⇒ match eq_b8 abs_magic (magic_of_opcode WAIT) with
+           [ true ⇒ TickSUSP ? s_tmp3 WAIT_MODE
+           | false ⇒ TickOK ? s_tmp3
+           ]]]]].
+
+definition tick ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+ let opt_info ≝ get_clk_desc m t s in
+ match opt_info with
+  (* e' il momento del fetch *)
+  [ None ⇒ match fetch m t s with
+   (* errore nel fetch/decode? riportato in output, nessun avanzamento *)
+   [ FetchERR err ⇒ TickERR ? s err
+   (* nessun errore nel fetch *)
+   | FetchOK fetch_info cur_pc ⇒ match fetch_info with
+    [ quadrupleT pseudo mode _ tot_clk ⇒
+      match eq_b8 〈x0,x1〉 tot_clk with
+       (* un solo clk, execute subito *)
+       [ true ⇒ tick_execute m t s pseudo mode cur_pc
+       (* piu' clk, execute rimandata *)
+       | false ⇒ TickOK ? (set_clk_desc m t s (Some ? (quintupleT ????? 〈x0,x1〉 pseudo mode tot_clk cur_pc)))
+       ]
+      ]
+    ]
+  (* il fetch e' gia' stato eseguito, e' il turno di execute? *)
+  | Some info ⇒ match info with [ quintupleT cur_clk pseudo mode tot_clk cur_pc ⇒ 
+   match eq_b8 (succ_b8 cur_clk) tot_clk with
+    (* si *)
+    [ true ⇒ tick_execute m t s pseudo mode cur_pc
+    (* no, avanzamento cur_clk *)
+    | false ⇒ TickOK ? (set_clk_desc m t s (Some ? (quintupleT ????? (succ_b8 cur_clk) pseudo mode tot_clk cur_pc)))
+    ]
+   ]
+  ].
+
+(* ********** *)
+(* ESECUZIONE *)
+(* ********** *)
+
+let rec execute (m:mcu_type) (t:memory_impl) (s:tick_result (any_status m t)) (n:nat) on n ≝
+ match s with
+  [ TickERR s' error ⇒ TickERR ? s' error
+  | TickSUSP s' susp ⇒ TickSUSP ? s' susp
+  | TickOK s'        ⇒ match n with [ O ⇒ TickOK ? s' | S n' ⇒ execute m t (tick m t s') n' ]
+  ].
+
+lemma breakpoint:
+ ∀m,u,s,n1,n2. execute m u s (n1 + n2) = execute m u (execute m u s n1) n2.
+ intros;
+ generalize in match s; clear s;
+ elim n1 0;
+  [ intros;
+    cases t;
+    reflexivity
+  | simplify;
+    intros (n K s);
+    cases s;
+     [ 1,2: simplify;
+       cases n2;
+       simplify;
+       reflexivity;
+       | simplify;
+         apply K;
+     ]
+  ]
+qed.
diff --git a/helm/software/matita/library/freescale/opcode.ma b/helm/software/matita/library/freescale/opcode.ma
new file mode 100644 (file)
index 0000000..27a49f1
--- /dev/null
@@ -0,0 +1,477 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/opcode/".
+
+(*include "/media/VIRTUOSO/freescale/aux_bases.ma".*)
+include "freescale/aux_bases.ma".
+
+(* ********************************************** *)
+(* MATTONI BASE PER DEFINIRE LE TABELLE DELLE MCU *)
+(* ********************************************** *)
+
+(* enumerazione delle ALU *)
+inductive mcu_type: Type ≝
+  HC05  : mcu_type
+| HC08  : mcu_type
+| HCS08 : mcu_type
+| RS08  : mcu_type.
+
+(* enumerazione delle modalita' di indirizzamento = caricamento degli operandi *)
+inductive instr_mode: Type ≝
+  (* INHERENT = nessun operando *)
+  MODE_INH  : instr_mode
+  (* INHERENT = nessun operando (A implicito) *)
+| MODE_INHA : instr_mode
+  (* INHERENT = nessun operando (X implicito) *)
+| MODE_INHX : instr_mode
+  (* INHERENT = nessun operando (H implicito) *)
+| MODE_INHH : instr_mode
+
+  (* IMMEDIATE = operando valore immediato byte = 0xbb *)
+| MODE_IMM1 : instr_mode
+  (* IMMEDIATE = operando valore immediato word = 0xwwww *)
+| MODE_IMM2 : instr_mode
+  (* DIRECT = operando offset byte = [0x00bb] *)
+| MODE_DIR1 : instr_mode
+  (* DIRECT = operando offset word = [0xwwww] *)
+| MODE_DIR2 : instr_mode
+  (* INDEXED = nessun operando (implicito [X] *)
+| MODE_IX0  : instr_mode
+  (* INDEXED = operando offset relativo byte = [X+0x00bb] *)
+| MODE_IX1  : instr_mode
+  (* INDEXED = operando offset relativo word = [X+0xwwww] *)
+| MODE_IX2  : instr_mode
+  (* INDEXED = operando offset relativo byte = [SP+0x00bb] *)
+| MODE_SP1  : instr_mode
+  (* INDEXED = operando offset relativo word = [SP+0xwwww] *)
+| MODE_SP2  : instr_mode
+
+  (* DIRECT → DIRECT = carica da diretto/scrive su diretto *)
+| MODE_DIR1_to_DIR1 : instr_mode
+  (* IMMEDIATE → DIRECT = carica da immediato/scrive su diretto *)
+| MODE_IMM1_to_DIR1 : instr_mode
+  (* INDEXED++ → DIRECT = carica da [X]/scrive su diretto/H:X++ *)
+| MODE_IX0p_to_DIR1 : instr_mode
+  (* DIRECT → INDEXED++ = carica da diretto/scrive su [X]/H:X++ *)
+| MODE_DIR1_to_IX0p : instr_mode
+
+  (* INHERENT(A) + IMMEDIATE *)
+| MODE_INHA_and_IMM1 : instr_mode
+  (* INHERENT(X) + IMMEDIATE *)
+| MODE_INHX_and_IMM1 : instr_mode
+  (* IMMEDIATE + IMMEDIATE *)
+| MODE_IMM1_and_IMM1 : instr_mode
+  (* DIRECT + IMMEDIATE *)
+| MODE_DIR1_and_IMM1 : instr_mode
+  (* INDEXED + IMMEDIATE *)
+| MODE_IX0_and_IMM1  : instr_mode
+  (* INDEXED++ + IMMEDIATE *)
+| MODE_IX0p_and_IMM1 : instr_mode
+  (* INDEXED + IMMEDIATE *)
+| MODE_IX1_and_IMM1  : instr_mode
+  (* INDEXED++ + IMMEDIATE *)
+| MODE_IX1p_and_IMM1 : instr_mode
+  (* INDEXED + IMMEDIATE *)
+| MODE_SP1_and_IMM1  : instr_mode
+
+  (* DIRECT(mTNY) = operando offset byte(maschera scrittura implicita 3 bit) *)
+  (* ex: DIR3 e' carica b, scrivi b con n-simo bit modificato *)
+| MODE_DIRn          : oct → instr_mode
+  (* DIRECT(mTNY) + IMMEDIATE = operando offset byte(maschera lettura implicita 3 bit) *)
+  (*                            + operando valore immediato byte  *)
+  (* ex: DIR2_and_IMM1 e' carica b, carica imm, restituisci n-simo bit di b + imm *)
+| MODE_DIRn_and_IMM1 : oct → instr_mode
+  (* TINY = nessun operando (diretto implicito 4bit = [0x00000000:0000iiii]) *)
+| MODE_TNY           : exadecim → instr_mode
+  (* SHORT = nessun operando (diretto implicito 5bit = [0x00000000:000iiiii]) *)
+| MODE_SRT           : bitrigesim → instr_mode
+.
+
+(* enumerazione delle istruzioni di tutte le ALU *)
+inductive opcode: Type ≝
+  ADC    : opcode (* add with carry *)
+| ADD    : opcode (* add *)
+| AIS    : opcode (* add immediate to SP *)
+| AIX    : opcode (* add immediate to X *)
+| AND    : opcode (* and *)
+| ASL    : opcode (* aritmetic shift left *)
+| ASR    : opcode (* aritmetic shift right *)
+| BCC    : opcode (* branch if C=0 *)
+| BCLRn  : opcode (* clear bit n *)
+| BCS    : opcode (* branch if C=1 *)
+| BEQ    : opcode (* branch if Z=1 *)
+| BGE    : opcode (* branch if N⊙V=0 (great or equal) *)
+| BGND   : opcode (* !!background mode!! *)
+| BGT    : opcode (* branch if Z|N⊙V=0 clear (great) *)
+| BHCC   : opcode (* branch if H=0 *)
+| BHCS   : opcode (* branch if H=1 *)
+| BHI    : opcode (* branch if C|Z=0, (higher) *)
+| BIH    : opcode (* branch if nIRQ=1 *)
+| BIL    : opcode (* branch if nIRQ=0 *)
+| BIT    : opcode (* flag = and (bit test) *)
+| BLE    : opcode (* branch if Z|N⊙V=1 (less or equal) *)
+| BLS    : opcode (* branch if C|Z=1 (lower or same) *)
+| BLT    : opcode (* branch if N⊙1=1 (less) *)
+| BMC    : opcode (* branch if I=0 (interrupt mask clear) *)
+| BMI    : opcode (* branch if N=1 (minus) *)
+| BMS    : opcode (* branch if I=1 (interrupt mask set) *)
+| BNE    : opcode (* branch if Z=0 *)
+| BPL    : opcode (* branch if N=0 (plus) *)
+| BRA    : opcode (* branch always *)
+| BRCLRn : opcode (* branch if bit n clear *)
+| BRN    : opcode (* branch never (nop) *)
+| BRSETn : opcode (* branch if bit n set *)
+| BSETn  : opcode (* set bit n *)
+| BSR    : opcode (* branch to subroutine *)
+| CBEQA  : opcode (* compare (A) and BEQ *)
+| CBEQX  : opcode (* compare (X) and BEQ *)
+| CLC    : opcode (* C=0 *)
+| CLI    : opcode (* I=0 *)
+| CLR    : opcode (* operand=0 *)
+| CMP    : opcode (* flag = sub (compare A) *)
+| COM    : opcode (* not (1 complement) *)
+| CPHX   : opcode (* flag = sub (compare H:X) *)
+| CPX    : opcode (* flag = sub (compare X) *)
+| DAA    : opcode (* decimal adjust A *)
+| DBNZ   : opcode (* dec and BNE *)
+| DEC    : opcode (* operand=operand-1 (decrement) *)
+| DIV    : opcode (* div *)
+| EOR    : opcode (* xor *)
+| INC    : opcode (* operand=operand+1 (increment) *)
+| JMP    : opcode (* jmp word [operand] *)
+| JSR    : opcode (* jmp to subroutine *)
+| LDA    : opcode (* load in A *)
+| LDHX   : opcode (* load in H:X *)
+| LDX    : opcode (* load in X *)
+| LSR    : opcode (* logical shift right *)
+| MOV    : opcode (* move *)
+| MUL    : opcode (* mul *)
+| NEG    : opcode (* neg (2 complement) *)
+| NOP    : opcode (* nop *)
+| NSA    : opcode (* nibble swap A (al:ah <- ah:al) *)
+| ORA    : opcode (* or *)
+| PSHA   : opcode (* push A *)
+| PSHH   : opcode (* push H *)
+| PSHX   : opcode (* push X *)
+| PULA   : opcode (* pop A *)
+| PULH   : opcode (* pop H *)
+| PULX   : opcode (* pop X *)
+| ROL    : opcode (* rotate left *)
+| ROR    : opcode (* rotate right *)
+| RSP    : opcode (* reset SP (0x00FF) *)
+| RTI    : opcode (* return from interrupt *)
+| RTS    : opcode (* return from subroutine *)
+| SBC    : opcode (* sub with carry*)
+| SEC    : opcode (* C=1 *)
+| SEI    : opcode (* I=1 *)
+| SHA    : opcode (* swap spc_high,A *)
+| SLA    : opcode (* swap spc_low,A *)
+| STA    : opcode (* store from A *)
+| STHX   : opcode (* store from H:X *)
+| STOP   : opcode (* !!stop mode!! *)
+| STX    : opcode (* store from X *)
+| SUB    : opcode (* sub *)
+| SWI    : opcode (* software interrupt *)
+| TAP    : opcode (* flag=A (transfer A to process status byte *)
+| TAX    : opcode (* X=A (transfer A to X) *)
+| TPA    : opcode (* A=flag (transfer process status byte to A) *)
+| TST    : opcode (* flag = sub (test) *)
+| TSX    : opcode (* X:H=SP (transfer SP to H:X) *)
+| TXA    : opcode (* A=X (transfer X to A) *)
+| TXS    : opcode (* SP=X:H (transfer H:X to SP) *)
+| WAIT   : opcode (* !!wait mode!! *)
+.
+
+(* introduzione di un tipo opcode dipendente dall'mcu_type (phantom type) *)
+inductive any_opcode (m:mcu_type) : Type ≝
+ anyOP : opcode → any_opcode m.
+
+coercion cic:/matita/freescale/opcode/any_opcode.ind#xpointer(1/1/1).
+
+(* raggruppamento di byte e word in un tipo unico *)
+inductive byte8_or_word16 : Type ≝
+  Byte: byte8  → byte8_or_word16
+| Word: word16 → byte8_or_word16.
+
+coercion cic:/matita/freescale/opcode/byte8_or_word16.ind#xpointer(1/1/1).
+coercion cic:/matita/freescale/opcode/byte8_or_word16.ind#xpointer(1/1/2).
+
+(* opcode → naturali, per usare eqb *)
+definition magic_of_opcode ≝
+λo:opcode.match o with
+[ ADC    ⇒ 〈x0,x0〉
+| ADD    ⇒ 〈x0,x1〉
+| AIS    ⇒ 〈x0,x2〉
+| AIX    ⇒ 〈x0,x3〉
+| AND    ⇒ 〈x0,x4〉
+| ASL    ⇒ 〈x0,x5〉
+| ASR    ⇒ 〈x0,x6〉
+| BCC    ⇒ 〈x0,x7〉
+| BCLRn  ⇒ 〈x0,x8〉
+| BCS    ⇒ 〈x0,x9〉
+| BEQ    ⇒ 〈x0,xA〉
+| BGE    ⇒ 〈x0,xB〉
+| BGND   ⇒ 〈x0,xC〉
+| BGT    ⇒ 〈x0,xD〉
+| BHCC   ⇒ 〈x0,xE〉
+| BHCS   ⇒ 〈x0,xF〉
+| BHI    ⇒ 〈x1,x0〉
+| BIH    ⇒ 〈x1,x1〉
+| BIL    ⇒ 〈x1,x2〉
+| BIT    ⇒ 〈x1,x3〉
+| BLE    ⇒ 〈x1,x4〉
+| BLS    ⇒ 〈x1,x5〉
+| BLT    ⇒ 〈x1,x6〉
+| BMC    ⇒ 〈x1,x7〉
+| BMI    ⇒ 〈x1,x8〉
+| BMS    ⇒ 〈x1,x9〉
+| BNE    ⇒ 〈x1,xA〉
+| BPL    ⇒ 〈x1,xB〉
+| BRA    ⇒ 〈x1,xC〉
+| BRCLRn ⇒ 〈x1,xD〉
+| BRN    ⇒ 〈x1,xE〉
+| BRSETn ⇒ 〈x1,xF〉
+| BSETn  ⇒ 〈x2,x0〉
+| BSR    ⇒ 〈x2,x1〉
+| CBEQA  ⇒ 〈x2,x2〉
+| CBEQX  ⇒ 〈x2,x3〉
+| CLC    ⇒ 〈x2,x4〉
+| CLI    ⇒ 〈x2,x5〉
+| CLR    ⇒ 〈x2,x6〉
+| CMP    ⇒ 〈x2,x7〉
+| COM    ⇒ 〈x2,x8〉
+| CPHX   ⇒ 〈x2,x9〉
+| CPX    ⇒ 〈x2,xA〉
+| DAA    ⇒ 〈x2,xB〉
+| DBNZ   ⇒ 〈x2,xC〉
+| DEC    ⇒ 〈x2,xD〉
+| DIV    ⇒ 〈x2,xE〉
+| EOR    ⇒ 〈x2,xF〉
+| INC    ⇒ 〈x3,x0〉
+| JMP    ⇒ 〈x3,x1〉
+| JSR    ⇒ 〈x3,x2〉
+| LDA    ⇒ 〈x3,x3〉
+| LDHX   ⇒ 〈x3,x4〉
+| LDX    ⇒ 〈x3,x5〉
+| LSR    ⇒ 〈x3,x6〉
+| MOV    ⇒ 〈x3,x7〉
+| MUL    ⇒ 〈x3,x8〉
+| NEG    ⇒ 〈x3,x9〉
+| NOP    ⇒ 〈x3,xA〉
+| NSA    ⇒ 〈x3,xB〉
+| ORA    ⇒ 〈x3,xC〉
+| PSHA   ⇒ 〈x3,xD〉
+| PSHH   ⇒ 〈x3,xE〉
+| PSHX   ⇒ 〈x3,xF〉
+| PULA   ⇒ 〈x4,x0〉
+| PULH   ⇒ 〈x4,x1〉
+| PULX   ⇒ 〈x4,x2〉
+| ROL    ⇒ 〈x4,x3〉
+| ROR    ⇒ 〈x4,x4〉
+| RSP    ⇒ 〈x4,x5〉
+| RTI    ⇒ 〈x4,x6〉
+| RTS    ⇒ 〈x4,x7〉
+| SBC    ⇒ 〈x4,x8〉
+| SEC    ⇒ 〈x4,x9〉
+| SEI    ⇒ 〈x4,xA〉
+| SHA    ⇒ 〈x4,xB〉
+| SLA    ⇒ 〈x4,xC〉
+| STA    ⇒ 〈x4,xD〉
+| STHX   ⇒ 〈x4,xE〉
+| STOP   ⇒ 〈x4,xF〉
+| STX    ⇒ 〈x5,x0〉
+| SUB    ⇒ 〈x5,x1〉
+| SWI    ⇒ 〈x5,x2〉
+| TAP    ⇒ 〈x5,x3〉
+| TAX    ⇒ 〈x5,x4〉
+| TPA    ⇒ 〈x5,x5〉
+| TST    ⇒ 〈x5,x6〉
+| TSX    ⇒ 〈x5,x7〉
+| TXA    ⇒ 〈x5,x8〉
+| TXS    ⇒ 〈x5,x9〉
+| WAIT   ⇒ 〈x5,xA〉
+].
+
+(* confronto fra opcode, legale solo se tipati sulla stessa mcu *)
+definition eqop ≝
+λm:mcu_type.λo:any_opcode m.λo':any_opcode m.match o with
+ [ anyOP p ⇒ match o' with
+  [ anyOP p' ⇒ (eq_b8 (magic_of_opcode p) (magic_of_opcode p')) ] ].
+
+(* instr_mode → naturali, per usare eqb *)
+definition magic_of_instr_mode ≝
+λi:instr_mode.match i with
+[ MODE_INH  ⇒ 〈x0,x0〉
+| MODE_INHA ⇒ 〈x0,x1〉
+| MODE_INHX ⇒ 〈x0,x2〉
+| MODE_INHH ⇒ 〈x0,x3〉
+| MODE_IMM1 ⇒ 〈x0,x4〉
+| MODE_IMM2 ⇒ 〈x0,x5〉
+| MODE_DIR1 ⇒ 〈x0,x6〉
+| MODE_DIR2 ⇒ 〈x0,x7〉
+| MODE_IX0  ⇒ 〈x0,x8〉
+| MODE_IX1  ⇒ 〈x0,x9〉
+| MODE_IX2  ⇒ 〈x0,xA〉
+| MODE_SP1  ⇒ 〈x0,xB〉
+| MODE_SP2  ⇒ 〈x0,xC〉
+
+| MODE_DIR1_to_DIR1 ⇒ 〈x0,xD〉
+| MODE_IMM1_to_DIR1 ⇒ 〈x0,xE〉
+| MODE_IX0p_to_DIR1 ⇒ 〈x0,xF〉
+| MODE_DIR1_to_IX0p ⇒ 〈x1,x0〉
+
+| MODE_INHA_and_IMM1 ⇒ 〈x1,x1〉
+| MODE_INHX_and_IMM1 ⇒ 〈x1,x2〉
+| MODE_IMM1_and_IMM1 ⇒ 〈x1,x3〉
+| MODE_DIR1_and_IMM1 ⇒ 〈x1,x4〉
+| MODE_IX0_and_IMM1  ⇒ 〈x1,x5〉
+| MODE_IX0p_and_IMM1 ⇒ 〈x1,x6〉
+| MODE_IX1_and_IMM1  ⇒ 〈x1,x7〉
+| MODE_IX1p_and_IMM1 ⇒ 〈x1,x8〉
+| MODE_SP1_and_IMM1  ⇒ 〈x1,x9〉
+
+  (* 26-33: bisogna considerare l'operando implicito *)
+| MODE_DIRn o          ⇒ plus_b8nc 〈x1,xA〉 〈x0,(exadecim_of_oct o)〉
+  (* 34-41: bisogna considerare l'operando implicito *)
+| MODE_DIRn_and_IMM1 o ⇒ plus_b8nc 〈x2,x2〉 〈x0,(exadecim_of_oct o)〉
+  (* 42-57: bisogna considerare l'operando implicito *)
+| MODE_TNY e           ⇒ plus_b8nc 〈x2,xA〉 〈x0,e〉
+  (* 58-99: bisogna considerare gli operandi impliciti *)
+| MODE_SRT t           ⇒ plus_b8nc 〈x3,xA〉 (byte8_of_bitrigesim t)
+].
+
+(* confronto fra instr_mode *)
+definition eqim ≝
+λi:instr_mode.λi':instr_mode.(eq_b8 (magic_of_instr_mode i) (magic_of_instr_mode i')).
+
+(* ********************************************* *)
+(* STRUMENTI PER LE DIMOSTRAZIONI DI CORRETTEZZA *)
+(* ********************************************* *)
+
+(* su tutta la lista quante volte compare il byte *)
+let rec get_byte_count (m:mcu_type) (b:byte8) (c:nat)
+ (l:list (Prod4T (any_opcode m) instr_mode byte8_or_word16 byte8)) on l ≝
+ match l with
+  [ nil ⇒ c
+  | cons hd tl ⇒ match thd4T ???? hd with
+   [ Byte b' ⇒ match eq_b8 b b' with
+    [ true ⇒ get_byte_count m b (S c) tl
+    | false ⇒ get_byte_count m b c tl
+    ]
+   | Word _ ⇒ get_byte_count m b c tl
+   ]
+  ].
+
+(* su tutta la lista quante volte compare la word (0x9E+byte) *)
+let rec get_word_count (m:mcu_type) (b:byte8) (c:nat)
+ (l:list (Prod4T (any_opcode m) instr_mode byte8_or_word16 byte8)) on l ≝
+ match l with
+  [ nil ⇒ c
+  | cons hd tl ⇒ match thd4T ???? hd with
+   [ Byte _ ⇒ get_word_count m b c tl
+   | Word w ⇒ match eq_w16 〈〈x9,xE〉:b〉 w with
+    [ true ⇒ get_word_count m b (S c) tl
+    | false ⇒ get_word_count m b c tl
+    ]
+   ]
+  ].
+
+(* b e' non implementato? *)
+let rec test_not_impl (b:byte8) (l:list byte8) on l ≝
+ match l with
+  [ nil ⇒ false
+  | cons hd tl ⇒ match eq_b8 b hd with
+   [ true ⇒ true
+   | false ⇒ test_not_impl b tl
+   ]
+  ].
+
+(* su tutta la lista quante volte compare la coppia opcode,instr_mode *)
+let rec get_OpIm_count (m:mcu_type) (o:any_opcode m) (i:instr_mode) (c:nat)
+ (l:list (Prod4T (any_opcode m) instr_mode byte8_or_word16 byte8)) on l ≝
+ match l with
+  [ nil ⇒ c
+  | cons hd tl ⇒
+   match (eqop m o (fst4T ???? hd)) ⊗
+         (eqim i (snd4T ???? hd)) with
+    [ true ⇒ get_OpIm_count m o i (S c) tl
+    | false ⇒ get_OpIm_count m o i c tl
+    ] 
+  ].
+
+(* iteratore sugli opcode *)
+definition forall_opcode ≝ λP.
+ P ADC    ⊗ P ADD    ⊗ P AIS    ⊗ P AIX    ⊗ P AND    ⊗ P ASL    ⊗ P ASR    ⊗ P BCC    ⊗
+ P BCLRn  ⊗ P BCS    ⊗ P BEQ    ⊗ P BGE    ⊗ P BGND   ⊗ P BGT    ⊗ P BHCC   ⊗ P BHCS   ⊗
+ P BHI    ⊗ P BIH    ⊗ P BIL    ⊗ P BIT    ⊗ P BLE    ⊗ P BLS    ⊗ P BLT    ⊗ P BMC    ⊗
+ P BMI    ⊗ P BMS    ⊗ P BNE    ⊗ P BPL    ⊗ P BRA    ⊗ P BRCLRn ⊗ P BRN    ⊗ P BRSETn ⊗
+ P BSETn  ⊗ P BSR    ⊗ P CBEQA  ⊗ P CBEQX  ⊗ P CLC    ⊗ P CLI    ⊗ P CLR    ⊗ P CMP    ⊗
+ P COM    ⊗ P CPHX   ⊗ P CPX    ⊗ P DAA    ⊗ P DBNZ   ⊗ P DEC    ⊗ P DIV    ⊗ P EOR    ⊗
+ P INC    ⊗ P JMP    ⊗ P JSR    ⊗ P LDA    ⊗ P LDHX   ⊗ P LDX    ⊗ P LSR    ⊗ P MOV    ⊗
+ P MUL    ⊗ P NEG    ⊗ P NOP    ⊗ P NSA    ⊗ P ORA    ⊗ P PSHA   ⊗ P PSHH   ⊗ P PSHX   ⊗
+ P PULA   ⊗ P PULH   ⊗ P PULX   ⊗ P ROL    ⊗ P ROR    ⊗ P RSP    ⊗ P RTI    ⊗ P RTS    ⊗
+ P SBC    ⊗ P SEC    ⊗ P SEI    ⊗ P SHA    ⊗ P SLA    ⊗ P STA    ⊗ P STHX   ⊗ P STOP   ⊗
+ P STX    ⊗ P SUB    ⊗ P SWI    ⊗ P TAP    ⊗ P TAX    ⊗ P TPA    ⊗ P TST    ⊗ P TSX    ⊗
+ P TXA    ⊗ P TXS    ⊗ P WAIT.
+
+(* iteratore sulle modalita' *)
+definition forall_instr_mode ≝ λP.
+  P MODE_INH
+⊗ P MODE_INHA 
+⊗ P MODE_INHX 
+⊗ P MODE_INHH
+
+⊗ P MODE_IMM1 
+⊗ P MODE_IMM2 
+⊗ P MODE_DIR1 
+⊗ P MODE_DIR2
+⊗ P MODE_IX0  
+⊗ P MODE_IX1  
+⊗ P MODE_IX2  
+⊗ P MODE_SP1  
+⊗ P MODE_SP2
+
+⊗ P MODE_DIR1_to_DIR1 
+⊗ P MODE_IMM1_to_DIR1
+⊗ P MODE_IX0p_to_DIR1 
+⊗ P MODE_DIR1_to_IX0p
+
+⊗ P MODE_INHA_and_IMM1 
+⊗ P MODE_INHX_and_IMM1 
+⊗ P MODE_IMM1_and_IMM1
+⊗ P MODE_DIR1_and_IMM1 
+⊗ P MODE_IX0_and_IMM1  
+⊗ P MODE_IX0p_and_IMM1
+⊗ P MODE_IX1_and_IMM1  
+⊗ P MODE_IX1p_and_IMM1 
+⊗ P MODE_SP1_and_IMM1
+
+⊗ forall_oct (λo. P (MODE_DIRn o))
+⊗ forall_oct (λo. P (MODE_DIRn_and_IMM1 o))
+⊗ forall_exadecim (λe. P (MODE_TNY e))
+⊗ forall_bitrigesim (λt. P (MODE_SRT t)).
diff --git a/helm/software/matita/library/freescale/status.ma b/helm/software/matita/library/freescale/status.ma
new file mode 100644 (file)
index 0000000..b2ca150
--- /dev/null
@@ -0,0 +1,999 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/status/".
+
+(*include "/media/VIRTUOSO/freescale/memory_abs.ma".*)
+include "freescale/memory_abs.ma".
+
+(* *********************************** *)
+(* STATUS INTERNO DEL PROCESSORE (ALU) *)
+(* *********************************** *)
+
+(* ALU dell'HC05 *)
+record alu_HC05: Type ≝
+ {
+ (* A: registo accumulatore *)
+ acc_low_reg_HC05 : byte8;
+ (* X: registro indice *)
+ indX_low_reg_HC05 : byte8;
+ (* SP: registo stack pointer *)
+ sp_reg_HC05 : word16;
+ (* modificatori di SP: per esempio e' definito come 0000000011xxxxxxb *)
+ (* la logica della sua costruzione e' quindi (SP∧mask)∨fix *)
+ (* totalmente racchiusa nella ALU, bastera' fare get(set(SP)) *)
+ sp_mask_HC05 : word16;
+ sp_fix_HC05 : word16;
+ (* PC: registro program counter *)
+ pc_reg_HC05 : word16;
+ (* modificatore di PC: per esempio e' definito come 00xxxxxxxxxxxxxxb *)
+ (* la logica della sua costruzione e' quindi (PC∧mask) *)
+ (* totalmente racchiusa nella ALU, bastera' fare get(set(PC)) *)
+ pc_mask_HC05 : word16;
+ (* H: flag semi-carry (somma nibble basso) *)
+ h_flag_HC05 : bool;
+ (* I: flag mascheramento degli interrupt mascherabili: 1=mascherati *)
+ i_flag_HC05 : bool;
+ (* N: flag segno/negativita' *)
+ n_flag_HC05 : bool;
+ (* Z: flag zero *)
+ z_flag_HC05 : bool;
+ (* C: flag carry *)
+ c_flag_HC05 : bool;
+ (* IRQ: flag che simula il pin esterno IRQ *)
+ irq_flag_HC05 : bool
+ }.
+
+notation "{hvbox('A_Reg' ≝ acclow ; break 'X_Reg' ≝ indxlow ; break 'Sp_Reg' ≝ sp ; break 'Sp_Mask' ≝ spm ; break 'Sp_Fix' ≝ spf ; break 'Pc_Reg' ≝ pc ; break 'Pc_Mask' ≝ pcm ; break 'H_Flag' ≝ hfl ; break 'I_Flag' ≝ ifl ; break 'N_Flag' ≝ nfl ; break 'Z_Flag' ≝ zfl ; break 'C_Flag' ≝ cfl ; break 'IRQ_Flag' ≝ irqfl)}"
+ non associative with precedence 80 for
+ @{ 'mk_alu_HC05 $acclow $indxlow $sp $spm $spf $pc $pcm $hfl $ifl $nfl $zfl $cfl $irqfl }.
+interpretation "mk_alu_HC05" 'mk_alu_HC05 acclow indxlow sp spm spf pc pcm hfl ifl nfl zfl cfl irqfl =
+ (cic:/matita/freescale/status/alu_HC05.ind#xpointer(1/1/1) acclow indxlow sp spm spf pc pcm hfl ifl nfl zfl cfl irqfl).
+(* ALU dell'HC08/HCS08 *) 
+record alu_HC08: Type ≝
+ {
+ (* A: registo accumulatore *)
+ acc_low_reg_HC08 : byte8;
+ (* X: registro indice parte bassa *)
+ indX_low_reg_HC08 : byte8;
+ (* H: registro indice parte alta *)
+ indX_high_reg_HC08 : byte8;
+ (* SP: registo stack pointer *)
+ sp_reg_HC08 : word16;
+ (* PC: registro program counter *)
+ pc_reg_HC08 : word16;
+ (* V: flag overflow *)
+ v_flag_HC08 : bool;
+ (* H: flag semi-carry (somma nibble basso) *)
+ h_flag_HC08 : bool;
+ (* I: flag mascheramento degli interrupt mascherabili: 1=mascherati *)
+ i_flag_HC08 : bool;
+ (* N: flag segno/negativita' *)
+ n_flag_HC08 : bool;
+ (* Z: flag zero *)
+ z_flag_HC08 : bool;
+ (* C: flag carry *)
+ c_flag_HC08 : bool;
+ (* IRQ: flag che simula il pin esterno IRQ *)
+ irq_flag_HC08 : bool
+ }.
+
+notation "{hvbox('A_Reg' ≝ acclow ; break 'X_Reg' ≝ indxlow ; break 'H_Reg' ≝ indxhigh ; break 'Sp_Reg' ≝ sp ; break 'Pc_Reg' ≝ pc ; break 'V_Flag' ≝ vfl ; break 'H_Flag' ≝ hfl ; break 'I_Flag' ≝ ifl ; break 'N_Flag' ≝ nfl ; break 'Z_Flag' ≝ zfl ; break 'C_Flag' ≝ cfl ; break 'IRQ_Flag' ≝ irqfl)}"
+ non associative with precedence 80 for
+ @{ 'mk_alu_HC08 $acclow $indxlow $indxhigh $sp $pc $vfl $hfl $ifl $nfl $zfl $cfl $irqfl }.
+interpretation "mk_alu_HC08" 'mk_alu_HC08 acclow indxlow indxhigh sp pc vfl hfl ifl nfl zfl cfl irqfl =
+ (cic:/matita/freescale/status/alu_HC08.ind#xpointer(1/1/1) acclow indxlow indxhigh sp pc vfl hfl ifl nfl zfl cfl irqfl).
+
+(* ALU dell'RS08 *)
+record alu_RS08: Type ≝
+ {
+ (* A: registo accumulatore *)
+ acc_low_reg_RS08 : byte8;
+ (* PC: registro program counter *)
+ pc_reg_RS08 : word16;
+ (* modificatore di PC: per esempio e' definito come 00xxxxxxxxxxxxxxb *)
+ (* la logica della sua costruzione e' quindi (PC∧mask) *)
+ (* totalmente racchiusa nella ALU, bastera' fare get(set(PC)) *)
+ pc_mask_RS08 : word16;
+ (* SPC: registro shadow program counter *)
+ (* NB: il suo modificatore e' lo stesso di PC *)
+ spc_reg_RS08 : word16;
+
+ (* X: registro indice parte bassa *)
+ (* NB: in realta' e' mappato in memoria e non risiede nella ALU *)
+ (* la lettura/scrittura avviene tramite la locazione [0x000F] *)
+ (* la funzione memory_filter_read/write si occupera' di intercettare *)
+ (* e deviare sul registro le letture/scritture (modulo load_write) *)
+ x_map_RS08 : byte8;
+ (* PS: registro selezione di pagina *)
+ (* serve a indirizzare la finestra RAM di 64b [0x00C0-0x00FF] *)
+ (* su tutta la memoria installata [0x0000-0x3FFF]: [00pp pppp ppxx xxxx] *)
+ (* NB: in realta' e' mappato in memoria e non risiede nella ALU *)
+ (* la lettura/scrittura avviene tramite la locazione [0x001F] *)
+ (* la funzione memory_filter_read/write si occupera' di intercettare *)
+ (* e deviare sul registro le letture/scritture (modulo load_write) *)
+ ps_map_RS08 : byte8;
+ (* Z: flag zero *)
+ z_flag_RS08 : bool;
+ (* C: flag carry *)
+ c_flag_RS08 : bool
+ }.
+
+notation "{hvbox('A_Reg' ≝ acclow ; break 'Pc_Reg' ≝ pc ; break 'Pc_Mask' ≝ pcm ; break 'Spc_Reg' ≝ spc ; break 'X_Map' ≝ xm ; break 'Ps_Map' ≝ psm ; break 'Z_Flag' ≝ zfl ; break 'C_Flag' ≝ cfl)}"
+ non associative with precedence 80 for
+ @{ 'mk_alu_RS08 $acclow $pc $pcm $spc $xm $psm $zfl $cfl }.
+interpretation "mk_alu_RS08" 'mk_alu_RS08 acclow pc pcm spc xm psm zfl cfl =
+ (cic:/matita/freescale/status/alu_RS08.ind#xpointer(1/1/1) acclow pc pcm spc xm psm zfl cfl).
+
+(* tipo processore dipendente dalla mcu, varia solo la ALU *)
+record any_status (mcu:mcu_type) (t:memory_impl): Type ≝
+ {
+ alu : match mcu with
+  [ HC05 ⇒ alu_HC05 | HC08 ⇒ alu_HC08 | HCS08 ⇒ alu_HC08 | RS08 ⇒ alu_RS08 ];
+
+ (* descritore della memoria *)
+ mem_desc : aux_mem_type t;
+ (* descrittore del tipo di memoria installata *)
+ chk_desc : aux_chk_type t;
+ (* descrittore del click = stato di avanzamento dell'esecuzione+cur_pc conseguente a fetch *)
+ (* 1) None = istruzione eseguita, attesa del fetch *)
+ (* 2) Some cur_clk,pseudo,mode,clks,cur_pc = fetch eseguito, countup a esecuzione *) 
+ clk_desc : option (Prod5T byte8 (any_opcode mcu) (instr_mode) byte8 word16)
+ }.
+
+(* evitare di mostrare la memoria/descrittore che impalla il visualizzatore *)
+notation > "{hvbox('Alu' ≝ alu ; break 'Clk' ≝ clk)}" non associative with precedence 80 
+ for @{ 'mk_any_status $alu $mem $chk $clk }.
+interpretation "mk_any_status" 'mk_any_status alu mem chk clk =
+ (cic:/matita/freescale/status/any_status.ind#xpointer(1/1/1) alu mem chk clk).
+
+(* **************** *)
+(* GETTER SPECIFICI *)
+(* **************** *)
+
+(* funzione ausiliaria per il tipaggio dei getter *)
+definition aux_get_typing ≝ λx:Type.λm:mcu_type.match m with
+ [ HC05 ⇒ alu_HC05 | HC08 ⇒ alu_HC08 | HCS08 ⇒ alu_HC08 | RS08 ⇒ alu_RS08 ] → x.
+
+(* REGISTRI *)
+
+(* getter di A, esiste sempre *)
+definition get_acc_8_low_reg ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+ match m
+  return aux_get_typing byte8 with
+ [ HC05 ⇒ acc_low_reg_HC05
+ | HC08 ⇒ acc_low_reg_HC08
+ | HCS08 ⇒ acc_low_reg_HC08
+ | RS08 ⇒ acc_low_reg_RS08 ]
+ (alu m t s).
+
+(* getter di X, non esiste sempre *)
+definition get_indX_8_low_reg ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+ match m
+  return aux_get_typing (option byte8) with 
+ [ HC05 ⇒ λalu.Some ? (indX_low_reg_HC05 alu)
+ | HC08 ⇒ λalu.Some ? (indX_low_reg_HC08 alu)
+ | HCS08 ⇒ λalu.Some ? (indX_low_reg_HC08 alu)
+ | RS08 ⇒ λalu.None ? ]
+ (alu m t s).
+
+(* getter di H, non esiste sempre *)
+definition get_indX_8_high_reg ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+ match m
+  return aux_get_typing (option byte8) with 
+ [ HC05 ⇒ λalu.None ?
+ | HC08 ⇒ λalu.Some ? (indX_high_reg_HC08 alu)
+ | HCS08 ⇒ λalu.Some ? (indX_high_reg_HC08 alu)
+ | RS08 ⇒ λalu.None ? ]
+ (alu m t s).
+
+(* getter di H:X, non esiste sempre *)
+definition get_indX_16_reg ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+ match m
+  return aux_get_typing (option word16) with 
+ [ HC05 ⇒ λalu.None ?
+ | HC08 ⇒ λalu.Some ? (mk_word16 (indX_high_reg_HC08 alu) (indX_low_reg_HC08 alu))
+ | HCS08 ⇒ λalu.Some ? (mk_word16 (indX_high_reg_HC08 alu) (indX_low_reg_HC08 alu))
+ | RS08 ⇒ λalu.None ? ]
+ (alu m t s).
+
+(* getter di SP, non esiste sempre *)
+definition get_sp_reg ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+ match m
+  return aux_get_typing (option word16) with 
+ [ HC05 ⇒ λalu.Some ? (sp_reg_HC05 alu)
+ | HC08 ⇒ λalu.Some ? (sp_reg_HC08 alu)
+ | HCS08 ⇒ λalu.Some ? (sp_reg_HC08 alu)
+ | RS08 ⇒ λalu.None ? ]
+ (alu m t s).
+
+(* getter di PC, esiste sempre *)
+definition get_pc_reg ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+ match m
+  return aux_get_typing word16 with 
+ [ HC05 ⇒ pc_reg_HC05
+ | HC08 ⇒ pc_reg_HC08
+ | HCS08 ⇒ pc_reg_HC08
+ | RS08 ⇒ pc_reg_RS08 ]
+ (alu m t s).
+
+(* getter di SPC, non esiste sempre *)
+definition get_spc_reg ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+ match m
+  return aux_get_typing (option word16) with 
+ [ HC05 ⇒ λalu.None ?
+ | HC08 ⇒ λalu.None ?
+ | HCS08 ⇒ λalu.None ?
+ | RS08 ⇒ λalu.Some ? (spc_reg_RS08 alu) ]
+ (alu m t s).
+
+(* REGISTRI MEMORY MAPPED *)
+
+(* getter di memory mapped X, non esiste sempre *)
+definition get_x_map ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+ match m
+  return aux_get_typing (option byte8) with 
+ [ HC05 ⇒ λalu.None ?
+ | HC08 ⇒ λalu.None ?
+ | HCS08 ⇒ λalu.None ?
+ | RS08 ⇒ λalu.Some ? (x_map_RS08 alu) ]
+ (alu m t s).
+
+(* getter di memory mapped PS, non esiste sempre *)
+definition get_ps_map ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+ match m
+  return aux_get_typing (option byte8) with 
+ [ HC05 ⇒ λalu.None ?
+ | HC08 ⇒ λalu.None ?
+ | HCS08 ⇒ λalu.None ?
+ | RS08 ⇒ λalu.Some ? (ps_map_RS08 alu) ]
+ (alu m t s).
+
+(* FLAG *)
+
+(* getter di V, non esiste sempre *)
+definition get_v_flag ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+ match m
+  return aux_get_typing (option bool) with 
+ [ HC05 ⇒ λalu.None ?
+ | HC08 ⇒ λalu.Some ? (v_flag_HC08 alu)
+ | HCS08 ⇒ λalu.Some ? (v_flag_HC08 alu)
+ | RS08 ⇒ λalu.None ? ]
+ (alu m t s).
+
+(* getter di H, non esiste sempre *)
+definition get_h_flag ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+ match m
+  return aux_get_typing (option bool) with 
+ [ HC05 ⇒ λalu.Some ? (h_flag_HC05 alu)
+ | HC08 ⇒ λalu.Some ? (h_flag_HC08 alu)
+ | HCS08 ⇒ λalu.Some ? (h_flag_HC08 alu)
+ | RS08 ⇒ λalu.None ? ]
+ (alu m t s).
+
+(* getter di I, non esiste sempre *)
+definition get_i_flag ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+ match m
+  return aux_get_typing (option bool) with 
+ [ HC05 ⇒ λalu.Some ? (i_flag_HC05 alu)
+ | HC08 ⇒ λalu.Some ? (i_flag_HC08 alu)
+ | HCS08 ⇒ λalu.Some ? (i_flag_HC08 alu)
+ | RS08 ⇒ λalu.None ? ]
+ (alu m t s).
+
+(* getter di N, non esiste sempre *)
+definition get_n_flag ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+ match m
+  return aux_get_typing (option bool) with 
+ [ HC05 ⇒ λalu.Some ? (n_flag_HC05 alu)
+ | HC08 ⇒ λalu.Some ? (n_flag_HC08 alu)
+ | HCS08 ⇒ λalu.Some ? (n_flag_HC08 alu)
+ | RS08 ⇒ λalu.None ? ]
+ (alu m t s).
+
+(* getter di Z, esiste sempre *)
+definition get_z_flag ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+ match m
+  return aux_get_typing bool with 
+ [ HC05 ⇒ z_flag_HC05
+ | HC08 ⇒ z_flag_HC08
+ | HCS08 ⇒ z_flag_HC08
+ | RS08 ⇒ z_flag_RS08 ]
+ (alu m t s).
+
+(* getter di C, esiste sempre *)
+definition get_c_flag ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+ match m
+  return aux_get_typing bool with 
+ [ HC05 ⇒ c_flag_HC05
+ | HC08 ⇒ c_flag_HC08
+ | HCS08 ⇒ c_flag_HC08
+ | RS08 ⇒ c_flag_RS08 ]
+ (alu m t s).
+
+(* getter di IRQ, non esiste sempre *)
+definition get_irq_flag ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+ match m
+  return aux_get_typing (option bool) with 
+ [ HC05 ⇒ λalu.Some ? (irq_flag_HC05 alu)
+ | HC08 ⇒ λalu.Some ? (irq_flag_HC08 alu)
+ | HCS08 ⇒ λalu.Some ? (irq_flag_HC08 alu)
+ | RS08 ⇒ λalu.None ? ]
+ (alu m t s).
+
+(* DESCRITTORI ESTERNI ALLA ALU *)
+
+(* getter della ALU, esiste sempre *)
+definition get_alu ≝ λm:mcu_type.λt:memory_impl.λs:any_status m t.alu m t s.
+
+(* getter della memoria, esiste sempre *)
+definition get_mem_desc ≝ λm:mcu_type.λt:memory_impl.λs:any_status m t.mem_desc m t s.
+
+(* getter del descrittore, esiste sempre *)
+definition get_chk_desc ≝ λm:mcu_type.λt:memory_impl.λs:any_status m t.chk_desc m t s.
+
+(* getter del clik, esiste sempre *)
+definition get_clk_desc ≝ λm:mcu_type.λt:memory_impl.λs:any_status m t.clk_desc m t s.
+
+(* ***************************** *)
+(* SETTER SPECIFICI FORTI/DEBOLI *)
+(* ***************************** *)
+
+(* funzione ausiliaria per il tipaggio dei setter forti *)
+definition aux_set_typing ≝ λx:Type.λm:mcu_type.
+  match m with [ HC05 ⇒ alu_HC05 | HC08 ⇒ alu_HC08 | HCS08 ⇒ alu_HC08 | RS08 ⇒ alu_RS08 ]
+  → x →
+  match m with [ HC05 ⇒ alu_HC05 | HC08 ⇒ alu_HC08 | HCS08 ⇒ alu_HC08 | RS08 ⇒ alu_RS08 ].
+
+(* funzione ausiliaria per il tipaggio dei setter deboli *)
+definition aux_set_typing_opt ≝ λx:Type.λm:mcu_type.option
+ (match m with [ HC05 ⇒ alu_HC05 | HC08 ⇒ alu_HC08 | HCS08 ⇒ alu_HC08 | RS08 ⇒ alu_RS08 ]
+  → x →
+  match m with [ HC05 ⇒ alu_HC05 | HC08 ⇒ alu_HC08 | HCS08 ⇒ alu_HC08 | RS08 ⇒ alu_RS08 ]).
+
+(* DESCRITTORI ESTERNI ALLA ALU *)
+
+(* setter forte della ALU *)
+definition set_alu ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λalu'.match s with
+ [ mk_any_status _ mem chk clk ⇒ mk_any_status m t alu' mem chk clk ].
+
+(* setter forte della memoria *)
+definition set_mem_desc ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λmem':aux_mem_type t.match s with
+ [ mk_any_status alu _ chk clk ⇒ mk_any_status m t alu mem' chk clk ].
+
+(* setter forte del descrittore *)
+definition set_chk_desc ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λchk':aux_chk_type t.match s with
+ [ mk_any_status alu mem _ clk ⇒ mk_any_status m t alu mem chk' clk ].
+
+(* setter forte del clik *)
+definition set_clk_desc ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.
+λclk':option (Prod5T byte8 (any_opcode m) (instr_mode) byte8 word16).match s with
+ [ mk_any_status alu mem chk _ ⇒ mk_any_status m t alu mem chk clk' ].
+
+(* REGISTRO A *)
+
+(* setter specifico HC05 di A *)
+definition set_acc_8_low_reg_HC05 ≝
+λalu.λacclow':byte8.match alu with 
+ [ mk_alu_HC05 _ indxlow sp spm spf pc pcm hfl ifl nfl zfl cfl irqfl ⇒
+  mk_alu_HC05 acclow' indxlow sp spm spf pc pcm hfl ifl nfl zfl cfl irqfl ].
+
+(* setter specifico HC08/HCS08 di A *)
+definition set_acc_8_low_reg_HC08 ≝
+λalu.λacclow':byte8.match alu with 
+ [ mk_alu_HC08 _ indxlow indxhigh sp pc vfl hfl ifl nfl zfl cfl irqfl ⇒
+  mk_alu_HC08 acclow' indxlow indxhigh sp pc vfl hfl ifl nfl zfl cfl irqfl ].
+
+(* setter specifico RS08 di A *)
+definition set_acc_8_low_reg_RS08 ≝ 
+λalu.λacclow':byte8.match alu with 
+ [ mk_alu_RS08 _ pc pcm spc xm psm zfl cfl ⇒
+  mk_alu_RS08 acclow' pc pcm spc xm psm zfl cfl ].
+
+(* setter forte di A *)
+definition set_acc_8_low_reg ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λacclow':byte8.
+ set_alu m t s 
+ (match m return aux_set_typing byte8 with
+  [ HC05 ⇒ set_acc_8_low_reg_HC05
+  | HC08 ⇒ set_acc_8_low_reg_HC08
+  | HCS08 ⇒ set_acc_8_low_reg_HC08
+  | RS08 ⇒ set_acc_8_low_reg_RS08
+  ] (alu m t s) acclow').
+
+(* REGISTRO X *)
+
+(* setter specifico HC05 di X *)
+definition set_indX_8_low_reg_HC05 ≝
+λalu.λindxlow':byte8.match alu with 
+ [ mk_alu_HC05 acclow _ sp spm spf pc pcm hfl ifl nfl zfl cfl irqfl ⇒
+  mk_alu_HC05 acclow indxlow' sp spm spf pc pcm hfl ifl nfl zfl cfl irqfl ].
+
+(* setter specifico HC08/HCS08 di X *)
+definition set_indX_8_low_reg_HC08 ≝
+λalu.λindxlow':byte8.match alu with 
+ [ mk_alu_HC08 acclow _ indxhigh sp pc vfl hfl ifl nfl zfl cfl irqfl ⇒
+  mk_alu_HC08 acclow indxlow' indxhigh sp pc vfl hfl ifl nfl zfl cfl irqfl ].
+
+(* setter forte di X *)
+definition set_indX_8_low_reg ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λindxlow':byte8.
+ opt_map ?? (match m return aux_set_typing_opt byte8 with
+             [ HC05 ⇒ Some ? set_indX_8_low_reg_HC05
+             | HC08 ⇒ Some ? set_indX_8_low_reg_HC08
+             | HCS08 ⇒ Some ? set_indX_8_low_reg_HC08
+             | RS08 ⇒ None ? ])
+ (λf.Some ? (set_alu m t s (f (alu m t s) indxlow'))).
+
+(* setter debole di X *)
+definition setweak_indX_8_low_reg ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λindxlow':byte8.
+ match set_indX_8_low_reg m t s indxlow' with
+  [ None ⇒ s | Some s' ⇒ s' ].
+
+(* REGISTRO H *)
+
+(* setter specifico HC08/HCS08 di H *)
+definition set_indX_8_high_reg_HC08 ≝
+λalu.λindxhigh':byte8.match alu with 
+ [ mk_alu_HC08 acclow indxlow _ sp pc vfl hfl ifl nfl zfl cfl irqfl ⇒
+  mk_alu_HC08 acclow indxlow indxhigh' sp pc vfl hfl ifl nfl zfl cfl irqfl ].
+
+(* setter forte di H *)
+definition set_indX_8_high_reg ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λindxhigh':byte8.
+ opt_map ?? (match m return aux_set_typing_opt byte8 with
+             [ HC05 ⇒ None ?
+             | HC08 ⇒ Some ? set_indX_8_high_reg_HC08
+             | HCS08 ⇒ Some ? set_indX_8_high_reg_HC08
+             | RS08 ⇒ None ? ])
+ (λf.Some ? (set_alu m t s (f (alu m t s) indxhigh'))).
+
+(* setter debole di H *)
+definition setweak_indX_8_high_reg ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λindxhigh':byte8.
+ match set_indX_8_high_reg m t s indxhigh' with
+  [ None ⇒ s | Some s' ⇒ s' ].
+
+(* REGISTRO H:X *)
+
+(* setter specifico HC08/HCS08 di H:X *)
+definition set_indX_16_reg_HC08 ≝
+λalu.λindx16:word16.match alu with 
+ [ mk_alu_HC08 acclow _ _ sp pc vfl hfl ifl nfl zfl cfl irqfl ⇒
+  mk_alu_HC08 acclow (w16l indx16) (w16h indx16) sp pc vfl hfl ifl nfl zfl cfl irqfl ].
+
+(* setter forte di H:X *)
+definition set_indX_16_reg ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λindx16:word16.
+ opt_map ?? (match m return aux_set_typing_opt word16 with
+             [ HC05 ⇒ None ?
+             | HC08 ⇒ Some ? set_indX_16_reg_HC08
+             | HCS08 ⇒ Some ? set_indX_16_reg_HC08
+             | RS08 ⇒ None ? ])
+ (λf.Some ? (set_alu m t s (f (alu m t s) indx16))).
+
+(* setter debole di H:X *)
+definition setweak_indX_16_reg ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λindx16:word16.
+ match set_indX_16_reg m t s indx16 with
+  [ None ⇒ s | Some s' ⇒ s' ].
+
+(* REGISTRO SP *)
+
+(* setter specifico HC05 di SP, effettua (SP∧mask)∨fix *)
+definition set_sp_reg_HC05 ≝
+λalu.λsp':word16.match alu with 
+ [ mk_alu_HC05 acclow indxlow _ spm spf pc pcm hfl ifl nfl zfl cfl irqfl ⇒
+  mk_alu_HC05 acclow indxlow (or_w16 (and_w16 sp' spm) spf) spm spf pc pcm hfl ifl nfl zfl cfl irqfl ].
+
+(* setter specifico HC08/HCS08 di SP *)
+definition set_sp_reg_HC08 ≝
+λalu.λsp':word16.match alu with 
+ [ mk_alu_HC08 acclow indxlow indxhigh _ pc vfl hfl ifl nfl zfl cfl irqfl ⇒
+  mk_alu_HC08 acclow indxlow indxhigh sp' pc vfl hfl ifl nfl zfl cfl irqfl ].
+
+(* setter forte di SP *)
+definition set_sp_reg ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λsp':word16.
+ opt_map ?? (match m return aux_set_typing_opt word16 with
+             [ HC05 ⇒ Some ? set_sp_reg_HC05
+             | HC08 ⇒ Some ? set_sp_reg_HC08
+             | HCS08 ⇒ Some ? set_sp_reg_HC08
+             | RS08 ⇒ None ? ])
+ (λf.Some ? (set_alu m t s (f (alu m t s) sp'))).
+
+(* setter debole di SP *)
+definition setweak_sp_reg ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λsp':word16.
+ match set_sp_reg m t s sp' with
+  [ None ⇒ s | Some s' ⇒ s' ].
+
+(* REGISTRO PC *)
+
+(* setter specifico HC05 di PC, effettua PC∧mask *)
+definition set_pc_reg_HC05 ≝
+λalu.λpc':word16.match alu with 
+ [ mk_alu_HC05 acclow indxlow sp spm spf _ pcm hfl ifl nfl zfl cfl irqfl ⇒
+  mk_alu_HC05 acclow indxlow sp spm spf (and_w16 pc' pcm) pcm hfl ifl nfl zfl cfl irqfl ].
+
+(* setter specifico HC08/HCS08 di PC *)
+definition set_pc_reg_HC08 ≝
+λalu.λpc':word16.match alu with 
+ [ mk_alu_HC08 acclow indxlow indxhigh sp _ vfl hfl ifl nfl zfl cfl irqfl ⇒
+  mk_alu_HC08 acclow indxlow indxhigh sp pc' vfl hfl ifl nfl zfl cfl irqfl ].
+
+(* setter specifico RS08 di PC, effettua PC∧mask *)
+definition set_pc_reg_RS08 ≝ 
+λalu.λpc':word16.match alu with 
+ [ mk_alu_RS08 acclow _ pcm spc xm psm zfl cfl ⇒
+  mk_alu_RS08 acclow (and_w16 pc' pcm) pcm spc xm psm zfl cfl ].
+
+(* setter forte di PC *)
+definition set_pc_reg ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λpc':word16.
+ set_alu m t s 
+ (match m return aux_set_typing word16 with
+  [ HC05 ⇒ set_pc_reg_HC05
+  | HC08 ⇒ set_pc_reg_HC08
+  | HCS08 ⇒ set_pc_reg_HC08
+  | RS08 ⇒ set_pc_reg_RS08
+  ] (alu m t s) pc').
+
+(* REGISTRO SPC *)
+
+(* setter specifico RS08 di SPC, effettua SPC∧mask *)
+definition set_spc_reg_RS08 ≝ 
+λalu.λspc':word16.match alu with 
+ [ mk_alu_RS08 acclow pc pcm _ xm psm zfl cfl ⇒
+  mk_alu_RS08 acclow pc pcm (and_w16 spc' pcm) xm psm zfl cfl ].
+
+(* setter forte di SPC *)
+definition set_spc_reg ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λspc':word16.
+ opt_map ?? (match m return aux_set_typing_opt word16 with
+             [ HC05 ⇒ None ?
+             | HC08 ⇒ None ?
+             | HCS08 ⇒ None ?
+             | RS08 ⇒ Some ? set_spc_reg_RS08 ])
+ (λf.Some ? (set_alu m t s (f (alu m t s) spc'))).
+
+(* setter debole di SPC *)
+definition setweak_spc_reg ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λspc':word16.
+ match set_spc_reg m t s spc' with
+  [ None ⇒ s | Some s' ⇒ s' ].
+
+(* REGISTRO MEMORY MAPPED X *)
+
+(* setter specifico RS08 di memory mapped X *)
+definition set_x_map_RS08 ≝ 
+λalu.λxm':byte8.match alu with 
+ [ mk_alu_RS08 acclow pc pcm spc _ psm zfl cfl ⇒
+  mk_alu_RS08 acclow pc pcm spc xm' psm zfl cfl ].
+
+(* setter forte di memory mapped X *)
+definition set_x_map ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λxm':byte8.
+ opt_map ?? (match m return aux_set_typing_opt byte8 with
+             [ HC05 ⇒ None ?
+             | HC08 ⇒ None ?
+             | HCS08 ⇒ None ?
+             | RS08 ⇒ Some ? set_x_map_RS08 ])
+ (λf.Some ? (set_alu m t s (f (alu m t s) xm'))).
+
+(* setter debole di memory mapped X *)
+definition setweak_x_map ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λxm':byte8.
+ match set_x_map m t s xm' with
+  [ None ⇒ s | Some s' ⇒ s' ].
+
+(* REGISTRO MEMORY MAPPED PS *)
+
+(* setter specifico RS08 di memory mapped PS *)
+definition set_ps_map_RS08 ≝ 
+λalu.λpsm':byte8.match alu with 
+ [ mk_alu_RS08 acclow pc pcm spc xm _ zfl cfl ⇒
+  mk_alu_RS08 acclow pc pcm spc xm psm' zfl cfl ].
+
+(* setter forte di memory mapped PS *)
+definition set_ps_map ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λpsm':byte8.
+ opt_map ?? (match m return aux_set_typing_opt byte8 with
+             [ HC05 ⇒ None ?
+             | HC08 ⇒ None ?
+             | HCS08 ⇒ None ?
+             | RS08 ⇒ Some ? set_ps_map_RS08 ])
+ (λf.Some ? (set_alu m t s (f (alu m t s) psm'))).
+
+(* setter debole di memory mapped PS *)
+definition setweak_ps_map ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λpsm':byte8.
+ match set_ps_map m t s psm' with
+  [ None ⇒ s | Some s' ⇒ s' ].
+
+(* FLAG V *)
+
+(* setter specifico HC08/HCS08 di V *)
+definition set_v_flag_HC08 ≝
+λalu.λvfl':bool.match alu with 
+ [ mk_alu_HC08 acclow indxlow indxhigh sp pc _ hfl ifl nfl zfl cfl irqfl ⇒
+  mk_alu_HC08 acclow indxlow indxhigh sp pc vfl' hfl ifl nfl zfl cfl irqfl ].
+
+(* setter forte di V *)
+definition set_v_flag ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λvfl':bool.
+ opt_map ?? (match m return aux_set_typing_opt bool with
+             [ HC05 ⇒ None ?
+             | HC08 ⇒ Some ? set_v_flag_HC08
+             | HCS08 ⇒ Some ? set_v_flag_HC08
+             | RS08 ⇒ None ? ])
+ (λf.Some ? (set_alu m t s (f (alu m t s) vfl'))).
+
+(* setter debole  di V *)
+definition setweak_v_flag ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λvfl':bool.
+ match set_v_flag m t s vfl' with
+  [ None ⇒ s | Some s' ⇒ s' ].
+
+(* FLAG H *)
+
+(* setter specifico HC05 di H *)
+definition set_h_flag_HC05 ≝
+λalu.λhfl':bool.match alu with 
+ [ mk_alu_HC05 acclow indxlow sp spm spf pc pcm _ ifl nfl zfl cfl irqfl ⇒
+  mk_alu_HC05 acclow indxlow sp spm spf pc pcm hfl' ifl nfl zfl cfl irqfl ].
+
+(* setter specifico HC08/HCS08 di H *)
+definition set_h_flag_HC08 ≝
+λalu.λhfl':bool.match alu with 
+ [ mk_alu_HC08 acclow indxlow indxhigh sp pc vfl _ ifl nfl zfl cfl irqfl ⇒
+  mk_alu_HC08 acclow indxlow indxhigh sp pc vfl hfl' ifl nfl zfl cfl irqfl ].
+
+(* setter forte di H *)
+definition set_h_flag ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λhfl':bool.
+ opt_map ?? (match m return aux_set_typing_opt bool with
+             [ HC05 ⇒ Some ? set_h_flag_HC05
+             | HC08 ⇒ Some ? set_h_flag_HC08
+             | HCS08 ⇒ Some ? set_h_flag_HC08
+             | RS08 ⇒ None ? ])
+ (λf.Some ? (set_alu m t s (f (alu m t s) hfl'))).
+
+(* setter debole di H *)
+definition setweak_h_flag ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λhfl':bool.
+ match set_h_flag m t s hfl' with
+  [ None ⇒ s | Some s' ⇒ s' ].
+
+(* FLAG I *)
+
+(* setter specifico HC05 di I *)
+definition set_i_flag_HC05 ≝
+λalu.λifl':bool.match alu with 
+ [ mk_alu_HC05 acclow indxlow sp spm spf pc pcm hfl _ nfl zfl cfl irqfl ⇒
+  mk_alu_HC05 acclow indxlow sp spm spf pc pcm hfl ifl' nfl zfl cfl irqfl ].
+
+(* setter specifico HC08/HCS08 di I *)
+definition set_i_flag_HC08 ≝
+λalu.λifl':bool.match alu with 
+ [ mk_alu_HC08 acclow indxlow indxhigh sp pc vfl hfl _ nfl zfl cfl irqfl ⇒
+  mk_alu_HC08 acclow indxlow indxhigh sp pc vfl hfl ifl' nfl zfl cfl irqfl ].
+
+(* setter forte di I *)
+definition set_i_flag ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λifl':bool.
+ opt_map ?? (match m return aux_set_typing_opt bool with
+             [ HC05 ⇒ Some ? set_i_flag_HC05
+             | HC08 ⇒ Some ? set_i_flag_HC08
+             | HCS08 ⇒ Some ? set_i_flag_HC08
+             | RS08 ⇒ None ? ])
+ (λf.Some ? (set_alu m t s (f (alu m t s) ifl'))).
+
+(* setter debole di I *)
+definition setweak_i_flag ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λifl':bool.
+ match set_i_flag m t s ifl' with
+  [ None ⇒ s | Some s' ⇒ s' ].
+
+(* FLAG N *)
+
+(* setter specifico HC05 di N *)
+definition set_n_flag_HC05 ≝
+λalu.λnfl':bool.match alu with 
+ [ mk_alu_HC05 acclow indxlow sp spm spf pc pcm hfl ifl _ zfl cfl irqfl ⇒
+  mk_alu_HC05 acclow indxlow sp spm spf pc pcm hfl ifl nfl' zfl cfl irqfl ].
+
+(* setter specifico HC08/HCS08 di N *)
+definition set_n_flag_HC08 ≝
+λalu.λnfl':bool.match alu with 
+ [ mk_alu_HC08 acclow indxlow indxhigh sp pc vfl hfl ifl _ zfl cfl irqfl ⇒
+  mk_alu_HC08 acclow indxlow indxhigh sp pc vfl hfl ifl nfl' zfl cfl irqfl ].
+
+(* setter forte di N *)
+definition set_n_flag ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λnfl':bool.
+ opt_map ?? (match m return aux_set_typing_opt bool with
+             [ HC05 ⇒ Some ? set_n_flag_HC05
+             | HC08 ⇒ Some ? set_n_flag_HC08
+             | HCS08 ⇒ Some ? set_n_flag_HC08
+             | RS08 ⇒ None ? ])
+ (λf.Some ? (set_alu m t s (f (alu m t s) nfl'))).
+
+(* setter debole di N *)
+definition setweak_n_flag ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λnfl':bool.
+ match set_n_flag m t s nfl' with
+  [ None ⇒ s | Some s' ⇒ s' ].
+
+(* FLAG Z *)
+
+(* setter specifico HC05 di Z *)
+definition set_z_flag_HC05 ≝
+λalu.λzfl':bool.match alu with 
+ [ mk_alu_HC05 acclow indxlow sp spm spf pc pcm hfl ifl nfl _ cfl irqfl ⇒
+  mk_alu_HC05 acclow indxlow sp spm spf pc pcm hfl ifl nfl zfl' cfl irqfl ].
+
+(* setter specifico HC08/HCS08 di Z *)
+definition set_z_flag_HC08 ≝
+λalu.λzfl':bool.match alu with 
+ [ mk_alu_HC08 acclow indxlow indxhigh sp pc vfl hfl ifl nfl _ cfl irqfl ⇒
+  mk_alu_HC08 acclow indxlow indxhigh sp pc vfl hfl ifl nfl zfl' cfl irqfl ].
+
+(* setter sprcifico RS08 di Z *)
+definition set_z_flag_RS08 ≝ 
+λalu.λzfl':bool.match alu with 
+ [ mk_alu_RS08 acclow pc pcm spc xm psm _ cfl ⇒
+  mk_alu_RS08 acclow pc pcm spc xm psm zfl' cfl ].
+
+(* setter forte di Z *)
+definition set_z_flag ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λzfl':bool.
+ set_alu m t s 
+ (match m return aux_set_typing bool with
+  [ HC05 ⇒ set_z_flag_HC05
+  | HC08 ⇒ set_z_flag_HC08
+  | HCS08 ⇒ set_z_flag_HC08
+  | RS08 ⇒ set_z_flag_RS08
+  ] (alu m t s) zfl').
+
+(* FLAG C *)
+
+(* setter specifico HC05 di C *)
+definition set_c_flag_HC05 ≝
+λalu.λcfl':bool.match alu with 
+ [ mk_alu_HC05 acclow indxlow sp spm spf pc pcm hfl ifl nfl zfl _ irqfl ⇒
+  mk_alu_HC05 acclow indxlow sp spm spf pc pcm hfl ifl nfl zfl cfl' irqfl ].
+
+(* setter specifico HC08/HCS08 di C *)
+definition set_c_flag_HC08 ≝
+λalu.λcfl':bool.match alu with 
+ [ mk_alu_HC08 acclow indxlow indxhigh sp pc vfl hfl ifl nfl zfl _ irqfl ⇒
+  mk_alu_HC08 acclow indxlow indxhigh sp pc vfl hfl ifl nfl zfl cfl' irqfl ].
+
+(* setter specifico RS08 di C *)
+definition set_c_flag_RS08 ≝ 
+λalu.λcfl':bool.match alu with 
+ [ mk_alu_RS08 acclow pc pcm spc xm psm zfl _ ⇒
+  mk_alu_RS08 acclow pc pcm spc xm psm zfl cfl' ].
+
+(* setter forte di C *)
+definition set_c_flag ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λcfl':bool.
+ set_alu m t s 
+ (match m return aux_set_typing bool with
+  [ HC05 ⇒ set_c_flag_HC05
+  | HC08 ⇒ set_c_flag_HC08
+  | HCS08 ⇒ set_c_flag_HC08
+  | RS08 ⇒ set_c_flag_RS08
+  ] (alu m t s) cfl').
+
+(* FLAG IRQ *)
+
+(* setter specifico HC05 di IRQ *)
+definition set_irq_flag_HC05 ≝
+λalu.λirqfl':bool.match alu with 
+ [ mk_alu_HC05 acclow indxlow sp spm spf pc pcm hfl ifl nfl zfl cfl _ ⇒
+  mk_alu_HC05 acclow indxlow sp spm spf pc pcm hfl ifl nfl zfl cfl irqfl' ].
+
+(* setter specifico HC08/HCS08 di IRQ *)
+definition set_irq_flag_HC08 ≝
+λalu.λirqfl':bool.match alu with 
+ [ mk_alu_HC08 acclow indxlow indxhigh sp pc vfl hfl ifl nfl zfl cfl _ ⇒
+  mk_alu_HC08 acclow indxlow indxhigh sp pc vfl hfl ifl nfl zfl cfl irqfl' ].
+
+(* setter forte di IRQ *)
+definition set_irq_flag ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λirqfl':bool.
+ opt_map ?? (match m return aux_set_typing_opt bool with
+             [ HC05 ⇒ Some ? set_irq_flag_HC05
+             | HC08 ⇒ Some ? set_irq_flag_HC08
+             | HCS08 ⇒ Some ? set_irq_flag_HC08
+             | RS08 ⇒ None ? ])
+ (λf.Some ? (set_alu m t s (f (alu m t s) irqfl'))).
+
+(* setter debole di IRQ *)
+definition setweak_irq_flag ≝
+λm:mcu_type.λt:memory_impl.λs:any_status m t.λirqfl':bool.
+ match set_irq_flag m t s irqfl' with
+  [ None ⇒ s | Some s' ⇒ s' ].
+
+(* ***************** *)
+(* CONFRONTO FRA ALU *)
+(* ***************** *)
+
+(* confronto registro per registro dell'HC05 *)
+definition eq_alu_HC05 ≝
+λalu1,alu2:alu_HC05.
+ match alu1 with
+  [ mk_alu_HC05 acclow1 indxlow1 sp1 spm1 spf1 pc1 pcm1 hfl1 ifl1 nfl1 zfl1 cfl1 irqfl1 ⇒
+ match alu2 with
+  [ mk_alu_HC05 acclow2 indxlow2 sp2 spm2 spf2 pc2 pcm2 hfl2 ifl2 nfl2 zfl2 cfl2 irqfl2 ⇒
+   (eq_b8 acclow1 acclow2) ⊗
+   (eq_b8 indxlow1 indxlow2) ⊗
+   (eq_w16 sp1 sp2) ⊗
+   (eq_w16 spm1 spm2) ⊗
+   (eq_w16 spf1 spf2) ⊗
+   (eq_w16 pc1 pc2) ⊗
+   (eq_w16 pcm1 pcm2) ⊗
+   (eq_bool hfl1 hfl2) ⊗
+   (eq_bool ifl1 ifl2) ⊗
+   (eq_bool nfl1 nfl2) ⊗
+   (eq_bool zfl1 zfl2) ⊗
+   (eq_bool cfl1 cfl2) ⊗
+   (eq_bool irqfl1 irqfl2) ]].
+
+(* confronto registro per registro dell'HC08 *)
+definition eq_alu_HC08 ≝
+λalu1,alu2:alu_HC08.
+ match alu1 with
+  [ mk_alu_HC08 acclow1 indxlow1 indxhigh1 sp1 pc1 vfl1 hfl1 ifl1 nfl1 zfl1 cfl1 irqfl1 ⇒
+ match alu2 with
+  [ mk_alu_HC08 acclow2 indxlow2 indxhigh2 sp2 pc2 vfl2 hfl2 ifl2 nfl2 zfl2 cfl2 irqfl2 ⇒
+   (eq_b8 acclow1 acclow2) ⊗
+   (eq_b8 indxlow1 indxlow2) ⊗
+   (eq_b8 indxhigh1 indxhigh2) ⊗
+   (eq_w16 sp1 sp2) ⊗
+   (eq_w16 pc1 pc2) ⊗
+   (eq_bool vfl1 vfl2) ⊗
+   (eq_bool hfl1 hfl2) ⊗
+   (eq_bool ifl1 ifl2) ⊗
+   (eq_bool nfl1 nfl2) ⊗
+   (eq_bool zfl1 zfl2) ⊗
+   (eq_bool cfl1 cfl2) ⊗
+   (eq_bool irqfl1 irqfl2) ]].
+
+(* confronto registro per registro dell'RS08 *)
+definition eq_alu_RS08 ≝
+λalu1,alu2:alu_RS08.
+ match alu1 with
+  [ mk_alu_RS08 acclow1 pc1 pcm1 spc1 xm1 psm1 zfl1 cfl1 ⇒
+ match alu2 with
+  [ mk_alu_RS08 acclow2 pc2 pcm2 spc2 xm2 psm2 zfl2 cfl2 ⇒
+   (eq_b8 acclow1 acclow2) ⊗
+   (eq_w16 pc1 pc2) ⊗
+   (eq_w16 pcm1 pcm2) ⊗
+   (eq_w16 spc1 spc2) ⊗
+   (eq_b8 xm1 xm2) ⊗
+   (eq_b8 psm1 psm2) ⊗
+   (eq_bool zfl1 zfl2) ⊗
+   (eq_bool cfl1 cfl2) ]].
+
+(* ******************** *)
+(* CONFRONTO FRA STATUS *)
+(* ******************** *)
+
+(* la lettura da memoria ritorna un option byte8, quindi bisogna poterli confrontare *)
+definition eq_b8_opt ≝
+λb1,b2:option byte8.
+ match b1 with
+  [ None ⇒ match b2 with
+   [ None ⇒ true | Some _ ⇒ false ]
+  | Some b1' ⇒ match b2 with
+   [ None ⇒ false | Some b2' ⇒ eq_b8 b1' b2' ]
+  ].
+
+(* confronto di una regione di memoria [inf,inf+n] *)
+let rec forall_memory_ranged
+ (t:memory_impl)
+ (chk1:aux_chk_type t) (chk2:aux_chk_type t)
+ (mem1:aux_mem_type t) (mem2:aux_mem_type t)
+ (inf:word16) (n:nat) on n ≝
+ match n return λ_.bool with
+  [ O ⇒ eq_b8_opt (mem_read t mem1 chk1 inf) (mem_read t mem2 chk2 inf)
+  | S n' ⇒ (eq_b8_opt (mem_read t mem1 chk1 (plus_w16nc inf (word16_of_nat n)))
+                      (mem_read t mem2 chk2 (plus_w16nc inf (word16_of_nat n)))) ⊗
+           (forall_memory_ranged t chk1 chk2 mem1 mem2 inf n')
+  ].
+
+(* il clk e' option (Prod5T byte8 (any_opcode m) (instr_mode) byte8 word16) *)
+definition eq_clk ≝
+λm:mcu_type.λc1,c2:option (Prod5T byte8 (any_opcode m) (instr_mode) byte8 word16).
+ match c1 with
+  [ None ⇒ match c2 with
+   [ None ⇒ true | Some _ ⇒ false ]
+  | Some c1' ⇒ match c2 with
+   [ None ⇒ false | Some c2' ⇒ (eq_b8 (fst5T ????? c1') (fst5T ????? c2')) ⊗
+                               (eqop m (snd5T ????? c1') (snd5T ????? c2')) ⊗
+                               (eqim (thd5T ????? c1') (thd5T ????? c2')) ⊗
+                               (eq_b8 (frth5T ????? c1') (frth5T ????? c2')) ⊗
+                               (eq_w16 (ffth5T ????? c1') (ffth5T ????? c2')) ]
+  ].
+
+(* generalizzazione del confronto fra stati *)
+definition eq_status ≝
+λm:mcu_type.λt:memory_impl.λs1,s2:any_status m t.λinf:word16.λn:nat.
+ match s1 with [ mk_any_status alu1 mem1 chk1 clk1 ⇒
+ match s2 with [ mk_any_status alu2 mem2 chk2 clk2 ⇒
+
+ (* 1) confronto della ALU *)
+ (match m return λm:mcu_type.
+   match m with
+    [ HC05 ⇒ alu_HC05 | HC08 ⇒ alu_HC08 | HCS08 ⇒ alu_HC08 | RS08 ⇒ alu_RS08 ] →
+   match m with
+    [ HC05 ⇒ alu_HC05 | HC08 ⇒ alu_HC08 | HCS08 ⇒ alu_HC08 | RS08 ⇒ alu_RS08 ] →
+   bool with
+  [ HC05 ⇒ eq_alu_HC05 | HC08 ⇒ eq_alu_HC08 | HCS08 ⇒ eq_alu_HC08 | RS08 ⇒ eq_alu_RS08 ]
+ alu1 alu2) ⊗
+
+ (* 2) confronto della memoria in [inf,inf+n] *)
+ (forall_memory_ranged t chk1 chk2 mem1 mem2 inf n) ⊗
+
+ (* 3) confronto del clik *)
+ (eq_clk m clk1 clk2)
+ ]].
+
+(* assioma da applicare per l'uguaglianza degli stati *)
+axiom eq_status_axiom :
+∀inf:word16.∀n:nat.∀m:mcu_type.∀t:memory_impl.∀s1,s2:any_status m t.
+ (eq_status m t s1 s2 inf n = true) →
+ s1 = s2.
diff --git a/helm/software/matita/library/freescale/table_HC05.ma b/helm/software/matita/library/freescale/table_HC05.ma
new file mode 100644 (file)
index 0000000..57a6bf4
--- /dev/null
@@ -0,0 +1,413 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/table_HC05/".
+
+(*include "/media/VIRTUOSO/freescale/opcode.ma".*)
+include "freescale/opcode.ma".
+
+(* ***************** *)
+(* TABELLA DELL'HC05 *)
+(* ***************** *)
+
+(* definizione come concatenazione finale di liste per velocizzare il parsing *)
+(* ogni riga e' (any_opcode m) (instr_mode) (opcode esadecimale) (#cicli esecuzione) *)
+(* NB: l'uso di any_opcode m + concatenazione finale tutte liste
+       impedisce di introdurre opcode disomogenei (per mcu) *)
+
+definition opcode_table_HC05_1 ≝
+[
+  quadrupleT ???? (anyOP HC05 ADC) MODE_IMM1 (Byte 〈xA,x9〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC05 ADC) MODE_DIR1 (Byte 〈xB,x9〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 ADC) MODE_DIR2 (Byte 〈xC,x9〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 ADC) MODE_IX2  (Byte 〈xD,x9〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 ADC) MODE_IX1  (Byte 〈xE,x9〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 ADC) MODE_IX0  (Byte 〈xF,x9〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC05_2 ≝
+[
+  quadrupleT ???? (anyOP HC05 ADD) MODE_IMM1 (Byte 〈xA,xB〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC05 ADD) MODE_DIR1 (Byte 〈xB,xB〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 ADD) MODE_DIR2 (Byte 〈xC,xB〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 ADD) MODE_IX2  (Byte 〈xD,xB〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 ADD) MODE_IX1  (Byte 〈xE,xB〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 ADD) MODE_IX0  (Byte 〈xF,xB〉) 〈x0,x3〉
+].
+
+definition opcode_table_HC05_3 ≝
+[
+  quadrupleT ???? (anyOP HC05 AND) MODE_IMM1 (Byte 〈xA,x4〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC05 AND) MODE_DIR1 (Byte 〈xB,x4〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 AND) MODE_DIR2 (Byte 〈xC,x4〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 AND) MODE_IX2  (Byte 〈xD,x4〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 AND) MODE_IX1  (Byte 〈xE,x4〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 AND) MODE_IX0  (Byte 〈xF,x4〉) 〈x0,x3〉
+].
+
+definition opcode_table_HC05_4 ≝
+[
+  quadrupleT ???? (anyOP HC05 ASL) MODE_DIR1 (Byte 〈x3,x8〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 ASL) MODE_INHA (Byte 〈x4,x8〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 ASL) MODE_INHX (Byte 〈x5,x8〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 ASL) MODE_IX1  (Byte 〈x6,x8〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HC05 ASL) MODE_IX0  (Byte 〈x7,x8〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC05_5 ≝
+[
+  quadrupleT ???? (anyOP HC05 ASR) MODE_DIR1 (Byte 〈x3,x7〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 ASR) MODE_INHA (Byte 〈x4,x7〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 ASR) MODE_INHX (Byte 〈x5,x7〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 ASR) MODE_IX1  (Byte 〈x6,x7〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HC05 ASR) MODE_IX0  (Byte 〈x7,x7〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC05_6 ≝
+[
+  quadrupleT ???? (anyOP HC05 BRA ) MODE_IMM1 (Byte 〈x2,x0〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 BRN ) MODE_IMM1 (Byte 〈x2,x1〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 BHI ) MODE_IMM1 (Byte 〈x2,x2〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 BLS ) MODE_IMM1 (Byte 〈x2,x3〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 BCC ) MODE_IMM1 (Byte 〈x2,x4〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 BCS ) MODE_IMM1 (Byte 〈x2,x5〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 BNE ) MODE_IMM1 (Byte 〈x2,x6〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 BEQ ) MODE_IMM1 (Byte 〈x2,x7〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 BHCC) MODE_IMM1 (Byte 〈x2,x8〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 BHCS) MODE_IMM1 (Byte 〈x2,x9〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 BPL ) MODE_IMM1 (Byte 〈x2,xA〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 BMI ) MODE_IMM1 (Byte 〈x2,xB〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 BMC ) MODE_IMM1 (Byte 〈x2,xC〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 BMS ) MODE_IMM1 (Byte 〈x2,xD〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 BIL ) MODE_IMM1 (Byte 〈x2,xE〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 BIH ) MODE_IMM1 (Byte 〈x2,xF〉) 〈x0,x3〉
+].
+
+definition opcode_table_HC05_7 ≝
+[
+  quadrupleT ???? (anyOP HC05 BSETn) (MODE_DIRn o0) (Byte 〈x1,x0〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BCLRn) (MODE_DIRn o0) (Byte 〈x1,x1〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BSETn) (MODE_DIRn o1) (Byte 〈x1,x2〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BCLRn) (MODE_DIRn o1) (Byte 〈x1,x3〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BSETn) (MODE_DIRn o2) (Byte 〈x1,x4〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BCLRn) (MODE_DIRn o2) (Byte 〈x1,x5〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BSETn) (MODE_DIRn o3) (Byte 〈x1,x6〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BCLRn) (MODE_DIRn o3) (Byte 〈x1,x7〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BSETn) (MODE_DIRn o4) (Byte 〈x1,x8〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BCLRn) (MODE_DIRn o4) (Byte 〈x1,x9〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BSETn) (MODE_DIRn o5) (Byte 〈x1,xA〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BCLRn) (MODE_DIRn o5) (Byte 〈x1,xB〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BSETn) (MODE_DIRn o6) (Byte 〈x1,xC〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BCLRn) (MODE_DIRn o6) (Byte 〈x1,xD〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BSETn) (MODE_DIRn o7) (Byte 〈x1,xE〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BCLRn) (MODE_DIRn o7) (Byte 〈x1,xF〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC05_8 ≝
+[
+  quadrupleT ???? (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x0〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x1〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x2〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x3〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x4〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x5〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x6〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x7〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x8〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x9〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xA〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xB〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xC〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xD〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BRSETn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xE〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BRCLRn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xF〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC05_9 ≝
+[
+  quadrupleT ???? (anyOP HC05 BIT) MODE_IMM1 (Byte 〈xA,x5〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC05 BIT) MODE_DIR1 (Byte 〈xB,x5〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 BIT) MODE_DIR2 (Byte 〈xC,x5〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 BIT) MODE_IX2  (Byte 〈xD,x5〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 BIT) MODE_IX1  (Byte 〈xE,x5〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 BIT) MODE_IX0  (Byte 〈xF,x5〉) 〈x0,x3〉
+].
+
+definition opcode_table_HC05_10 ≝
+[
+  quadrupleT ???? (anyOP HC05 MUL ) MODE_INH (Byte 〈x4,x2〉) 〈x0,xB〉
+; quadrupleT ???? (anyOP HC05 RTI ) MODE_INH (Byte 〈x8,x0〉) 〈x0,x9〉
+; quadrupleT ???? (anyOP HC05 RTS ) MODE_INH (Byte 〈x8,x1〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HC05 SWI ) MODE_INH (Byte 〈x8,x3〉) 〈x0,xA〉
+; quadrupleT ???? (anyOP HC05 STOP) MODE_INH (Byte 〈x8,xE〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC05 WAIT) MODE_INH (Byte 〈x8,xF〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC05 TAX ) MODE_INH (Byte 〈x9,x7〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC05 CLC ) MODE_INH (Byte 〈x9,x8〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC05 SEC ) MODE_INH (Byte 〈x9,x9〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC05 CLI ) MODE_INH (Byte 〈x9,xA〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC05 SEI ) MODE_INH (Byte 〈x9,xB〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC05 RSP ) MODE_INH (Byte 〈x9,xC〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC05 NOP ) MODE_INH (Byte 〈x9,xD〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC05 TXA ) MODE_INH (Byte 〈x9,xF〉) 〈x0,x2〉
+].
+
+definition opcode_table_HC05_11 ≝
+[
+  quadrupleT ???? (anyOP HC05 CLR) MODE_DIR1 (Byte 〈x3,xF〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 CLR) MODE_INHA (Byte 〈x4,xF〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 CLR) MODE_INHX (Byte 〈x5,xF〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 CLR) MODE_IX1  (Byte 〈x6,xF〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HC05 CLR) MODE_IX0  (Byte 〈x7,xF〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC05_12 ≝
+[
+  quadrupleT ???? (anyOP HC05 CMP) MODE_IMM1 (Byte 〈xA,x1〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC05 CMP) MODE_DIR1 (Byte 〈xB,x1〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 CMP) MODE_DIR2 (Byte 〈xC,x1〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 CMP) MODE_IX2  (Byte 〈xD,x1〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 CMP) MODE_IX1  (Byte 〈xE,x1〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 CMP) MODE_IX0  (Byte 〈xF,x1〉) 〈x0,x3〉
+].
+
+definition opcode_table_HC05_13 ≝
+[
+  quadrupleT ???? (anyOP HC05 COM) MODE_DIR1 (Byte 〈x3,x3〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 COM) MODE_INHA (Byte 〈x4,x3〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 COM) MODE_INHX (Byte 〈x5,x3〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 COM) MODE_IX1  (Byte 〈x6,x3〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HC05 COM) MODE_IX0  (Byte 〈x7,x3〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC05_14 ≝
+[
+  quadrupleT ???? (anyOP HC05 CPX) MODE_IMM1 (Byte 〈xA,x3〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC05 CPX) MODE_DIR1 (Byte 〈xB,x3〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 CPX) MODE_DIR2 (Byte 〈xC,x3〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 CPX) MODE_IX2  (Byte 〈xD,x3〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 CPX) MODE_IX1  (Byte 〈xE,x3〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 CPX) MODE_IX0  (Byte 〈xF,x3〉) 〈x0,x3〉
+].
+
+definition opcode_table_HC05_15 ≝
+[
+  quadrupleT ???? (anyOP HC05 DEC) MODE_DIR1 (Byte 〈x3,xA〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 DEC) MODE_INHA (Byte 〈x4,xA〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 DEC) MODE_INHX (Byte 〈x5,xA〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 DEC) MODE_IX1  (Byte 〈x6,xA〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HC05 DEC) MODE_IX0  (Byte 〈x7,xA〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC05_16 ≝
+[
+  quadrupleT ???? (anyOP HC05 EOR) MODE_IMM1 (Byte 〈xA,x8〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC05 EOR) MODE_DIR1 (Byte 〈xB,x8〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 EOR) MODE_DIR2 (Byte 〈xC,x8〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 EOR) MODE_IX2  (Byte 〈xD,x8〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 EOR) MODE_IX1  (Byte 〈xE,x8〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 EOR) MODE_IX0  (Byte 〈xF,x8〉) 〈x0,x3〉
+].
+
+definition opcode_table_HC05_17 ≝
+[
+  quadrupleT ???? (anyOP HC05 INC) MODE_DIR1 (Byte 〈x3,xC〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 INC) MODE_INHA (Byte 〈x4,xC〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 INC) MODE_INHX (Byte 〈x5,xC〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 INC) MODE_IX1  (Byte 〈x6,xC〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HC05 INC) MODE_IX0  (Byte 〈x7,xC〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC05_18 ≝
+[
+  quadrupleT ???? (anyOP HC05 JMP) MODE_DIR1 (Byte 〈xB,xC〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC05 JMP) MODE_DIR2 (Byte 〈xC,xC〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 JMP) MODE_IX2  (Byte 〈xD,xC〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 JMP) MODE_IX1  (Byte 〈xE,xC〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 JMP) MODE_IX0  (Byte 〈xF,xC〉) 〈x0,x2〉
+].
+
+definition opcode_table_HC05_19 ≝
+[
+  quadrupleT ???? (anyOP HC05 BSR) MODE_IMM1 (Byte 〈xA,xD〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HC05 JSR) MODE_DIR1 (Byte 〈xB,xD〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 JSR) MODE_DIR2 (Byte 〈xC,xD〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HC05 JSR) MODE_IX2  (Byte 〈xD,xD〉) 〈x0,x7〉
+; quadrupleT ???? (anyOP HC05 JSR) MODE_IX1  (Byte 〈xE,xD〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HC05 JSR) MODE_IX0  (Byte 〈xF,xD〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC05_20 ≝
+[
+  quadrupleT ???? (anyOP HC05 LDA) MODE_IMM1 (Byte 〈xA,x6〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC05 LDA) MODE_DIR1 (Byte 〈xB,x6〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 LDA) MODE_DIR2 (Byte 〈xC,x6〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 LDA) MODE_IX2  (Byte 〈xD,x6〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 LDA) MODE_IX1  (Byte 〈xE,x6〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 LDA) MODE_IX0  (Byte 〈xF,x6〉) 〈x0,x3〉
+].
+
+definition opcode_table_HC05_21 ≝
+[
+  quadrupleT ???? (anyOP HC05 LDX) MODE_IMM1 (Byte 〈xA,xE〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC05 LDX) MODE_DIR1 (Byte 〈xB,xE〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 LDX) MODE_DIR2 (Byte 〈xC,xE〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 LDX) MODE_IX2  (Byte 〈xD,xE〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 LDX) MODE_IX1  (Byte 〈xE,xE〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 LDX) MODE_IX0  (Byte 〈xF,xE〉) 〈x0,x3〉
+].
+
+definition opcode_table_HC05_22 ≝
+[
+  quadrupleT ???? (anyOP HC05 LSR) MODE_DIR1 (Byte 〈x3,x4〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 LSR) MODE_INHA (Byte 〈x4,x4〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 LSR) MODE_INHX (Byte 〈x5,x4〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 LSR) MODE_IX1  (Byte 〈x6,x4〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HC05 LSR) MODE_IX0  (Byte 〈x7,x4〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC05_23 ≝
+[
+  quadrupleT ???? (anyOP HC05 NEG) MODE_DIR1 (Byte 〈x3,x0〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 NEG) MODE_INHA (Byte 〈x4,x0〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 NEG) MODE_INHX (Byte 〈x5,x0〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 NEG) MODE_IX1  (Byte 〈x6,x0〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HC05 NEG) MODE_IX0  (Byte 〈x7,x0〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC05_24 ≝
+[
+  quadrupleT ???? (anyOP HC05 ORA) MODE_IMM1 (Byte 〈xA,xA〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC05 ORA) MODE_DIR1 (Byte 〈xB,xA〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 ORA) MODE_DIR2 (Byte 〈xC,xA〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 ORA) MODE_IX2  (Byte 〈xD,xA〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 ORA) MODE_IX1  (Byte 〈xE,xA〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 ORA) MODE_IX0  (Byte 〈xF,xA〉) 〈x0,x3〉
+].
+
+definition opcode_table_HC05_25 ≝
+[
+  quadrupleT ???? (anyOP HC05 ROL) MODE_DIR1 (Byte 〈x3,x9〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 ROL) MODE_INHA (Byte 〈x4,x9〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 ROL) MODE_INHX (Byte 〈x5,x9〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 ROL) MODE_IX1  (Byte 〈x6,x9〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HC05 ROL) MODE_IX0  (Byte 〈x7,x9〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC05_26 ≝
+[
+  quadrupleT ???? (anyOP HC05 ROR) MODE_DIR1 (Byte 〈x3,x6〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 ROR) MODE_INHA (Byte 〈x4,x6〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 ROR) MODE_INHX (Byte 〈x5,x6〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 ROR) MODE_IX1  (Byte 〈x6,x6〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HC05 ROR) MODE_IX0  (Byte 〈x7,x6〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC05_27 ≝
+[
+  quadrupleT ???? (anyOP HC05 SBC) MODE_IMM1 (Byte 〈xA,x2〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC05 SBC) MODE_DIR1 (Byte 〈xB,x2〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 SBC) MODE_DIR2 (Byte 〈xC,x2〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 SBC) MODE_IX2  (Byte 〈xD,x2〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 SBC) MODE_IX1  (Byte 〈xE,x2〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 SBC) MODE_IX0  (Byte 〈xF,x2〉) 〈x0,x3〉
+].
+
+definition opcode_table_HC05_28 ≝
+[
+  quadrupleT ???? (anyOP HC05 STA) MODE_DIR1 (Byte 〈xB,x7〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 STA) MODE_DIR2 (Byte 〈xC,x7〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 STA) MODE_IX2  (Byte 〈xD,x7〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HC05 STA) MODE_IX1  (Byte 〈xE,x7〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 STA) MODE_IX0  (Byte 〈xF,x7〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC05_29 ≝
+[
+  quadrupleT ???? (anyOP HC05 STX) MODE_DIR1 (Byte 〈xB,xF〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 STX) MODE_DIR2 (Byte 〈xC,xF〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 STX) MODE_IX2  (Byte 〈xD,xF〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HC05 STX) MODE_IX1  (Byte 〈xE,xF〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 STX) MODE_IX0  (Byte 〈xF,xF〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC05_30 ≝
+[
+  quadrupleT ???? (anyOP HC05 SUB) MODE_IMM1 (Byte 〈xA,x0〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC05 SUB) MODE_DIR1 (Byte 〈xB,x0〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 SUB) MODE_DIR2 (Byte 〈xC,x0〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 SUB) MODE_IX2  (Byte 〈xD,x0〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 SUB) MODE_IX1  (Byte 〈xE,x0〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 SUB) MODE_IX0  (Byte 〈xF,x0〉) 〈x0,x3〉
+].
+
+definition opcode_table_HC05_31 ≝
+[
+  quadrupleT ???? (anyOP HC05 TST) MODE_DIR1 (Byte 〈x3,xD〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC05 TST) MODE_INHA (Byte 〈x4,xD〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 TST) MODE_INHX (Byte 〈x5,xD〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC05 TST) MODE_IX1  (Byte 〈x6,xD〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC05 TST) MODE_IX0  (Byte 〈x7,xD〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC05 ≝
+ opcode_table_HC05_1  @ opcode_table_HC05_2  @ opcode_table_HC05_3  @ opcode_table_HC05_4  @
+ opcode_table_HC05_5  @ opcode_table_HC05_6  @ opcode_table_HC05_7  @ opcode_table_HC05_8  @
+ opcode_table_HC05_9  @ opcode_table_HC05_10 @ opcode_table_HC05_11 @ opcode_table_HC05_12 @
+ opcode_table_HC05_13 @ opcode_table_HC05_14 @ opcode_table_HC05_15 @ opcode_table_HC05_16 @
+ opcode_table_HC05_17 @ opcode_table_HC05_18 @ opcode_table_HC05_19 @ opcode_table_HC05_20 @
+ opcode_table_HC05_21 @ opcode_table_HC05_22 @ opcode_table_HC05_23 @ opcode_table_HC05_24 @
+ opcode_table_HC05_25 @ opcode_table_HC05_26 @ opcode_table_HC05_27 @ opcode_table_HC05_28 @
+ opcode_table_HC05_29 @ opcode_table_HC05_30 @ opcode_table_HC05_31.
+
+(* CORRETTEZZA *)
+
+(* HC05: opcode non implementati come da manuale *)
+definition HC05_not_impl ≝
+ [〈x3,x1〉;〈x3,x2〉;〈x3,x5〉;〈x3,xB〉;〈x3,xE〉
+ ;〈x4,x1〉;〈x4,x5〉;〈x4,xB〉;〈x4,xE〉
+ ;〈x5,x1〉;〈x5,x2〉;〈x5,x5〉;〈x5,xB〉;〈x5,xE〉
+ ;〈x6,x1〉;〈x6,x2〉;〈x6,x5〉;〈x6,xB〉;〈x6,xE〉
+ ;〈x7,x1〉;〈x7,x2〉;〈x7,x5〉;〈x7,xB〉;〈x7,xE〉
+ ;〈x8,x2〉;〈x8,x4〉;〈x8,x5〉;〈x8,x6〉;〈x8,x7〉;〈x8,x8〉;〈x8,x9〉;〈x8,xA〉;〈x8,xB〉;〈x8,xC〉;〈x8,xD〉
+ ;〈x9,x0〉;〈x9,x1〉;〈x9,x2〉;〈x9,x3〉;〈x9,x4〉;〈x9,x5〉;〈x9,x6〉;〈x9,xE〉
+ ;〈xA,x7〉;〈xA,xC〉;〈xA,xF〉
+ ].
+
+lemma ok_byte_table_HC05 : forall_byte8 (λb.
+ (test_not_impl b HC05_not_impl     ⊕ eqb (get_byte_count HC05 b 0 opcode_table_HC05) 1) ⊗
+ (⊖ (test_not_impl b HC05_not_impl) ⊕ eqb (get_byte_count HC05 b 0 opcode_table_HC05) 0))
+ = true.
+ reflexivity.
+qed.
+
+(*
+lemma ok_OpIm_table_HC05 :
+ forall_instr_mode (λi:instr_mode.
+ forall_opcode     (λop:opcode.
+  leb (get_OpIm_count HC05 (anyOP HC05 op) i 0 opcode_table_HC05) 1)) = true.
+ reflexivity.
+qed.
+*)
diff --git a/helm/software/matita/library/freescale/table_HC08.ma b/helm/software/matita/library/freescale/table_HC08.ma
new file mode 100644 (file)
index 0000000..61f7f3e
--- /dev/null
@@ -0,0 +1,546 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/table_HC08/".
+
+(*include "/media/VIRTUOSO/freescale/opcode.ma".*)
+include "freescale/opcode.ma".
+
+(* ***************** *)
+(* TABELLA DELL'HC08 *)
+(* ***************** *)
+
+(* definizione come concatenazione finale di liste per velocizzare il parsing *)
+(* ogni riga e' (any_opcode m) (instr_mode) (opcode esadecimale) (#cicli esecuzione) *)
+(* NB: l'uso di any_opcode m + concatenazione finale tutte liste
+       impedisce di introdurre opcode disomogenei (per mcu) *)
+
+definition opcode_table_HC08_1 ≝
+[
+  quadrupleT ???? (anyOP HC08 ADC) MODE_IMM1 (Byte 〈xA,x9〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 ADC) MODE_DIR1 (Byte 〈xB,x9〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 ADC) MODE_DIR2 (Byte 〈xC,x9〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 ADC) MODE_IX2  (Byte 〈xD,x9〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 ADC) MODE_IX1  (Byte 〈xE,x9〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 ADC) MODE_IX0  (Byte 〈xF,x9〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 ADC) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,x9〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 ADC) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,x9〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC08_2 ≝
+[
+  quadrupleT ???? (anyOP HC08 ADD) MODE_IMM1 (Byte 〈xA,xB〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 ADD) MODE_DIR1 (Byte 〈xB,xB〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 ADD) MODE_DIR2 (Byte 〈xC,xB〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 ADD) MODE_IX2  (Byte 〈xD,xB〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 ADD) MODE_IX1  (Byte 〈xE,xB〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 ADD) MODE_IX0  (Byte 〈xF,xB〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 ADD) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,xB〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 ADD) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,xB〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC08_3 ≝
+[
+  quadrupleT ???? (anyOP HC08 AND) MODE_IMM1 (Byte 〈xA,x4〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 AND) MODE_DIR1 (Byte 〈xB,x4〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 AND) MODE_DIR2 (Byte 〈xC,x4〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 AND) MODE_IX2  (Byte 〈xD,x4〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 AND) MODE_IX1  (Byte 〈xE,x4〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 AND) MODE_IX0  (Byte 〈xF,x4〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 AND) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,x4〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 AND) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,x4〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC08_4 ≝
+[
+  quadrupleT ???? (anyOP HC08 ASL) MODE_DIR1 (Byte 〈x3,x8〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 ASL) MODE_INHA (Byte 〈x4,x8〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 ASL) MODE_INHX (Byte 〈x5,x8〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 ASL) MODE_IX1  (Byte 〈x6,x8〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 ASL) MODE_IX0  (Byte 〈x7,x8〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 ASL) MODE_SP1  (Word 〈〈x9,xE〉:〈x6,x8〉〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC08_5 ≝
+[
+  quadrupleT ???? (anyOP HC08 ASR) MODE_DIR1 (Byte 〈x3,x7〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 ASR) MODE_INHA (Byte 〈x4,x7〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 ASR) MODE_INHX (Byte 〈x5,x7〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 ASR) MODE_IX1  (Byte 〈x6,x7〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 ASR) MODE_IX0  (Byte 〈x7,x7〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 ASR) MODE_SP1  (Word 〈〈x9,xE〉:〈x6,x7〉〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC08_6 ≝
+[
+  quadrupleT ???? (anyOP HC08 BRA ) MODE_IMM1 (Byte 〈x2,x0〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 BRN ) MODE_IMM1 (Byte 〈x2,x1〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 BHI ) MODE_IMM1 (Byte 〈x2,x2〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 BLS ) MODE_IMM1 (Byte 〈x2,x3〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 BCC ) MODE_IMM1 (Byte 〈x2,x4〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 BCS ) MODE_IMM1 (Byte 〈x2,x5〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 BNE ) MODE_IMM1 (Byte 〈x2,x6〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 BEQ ) MODE_IMM1 (Byte 〈x2,x7〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 BHCC) MODE_IMM1 (Byte 〈x2,x8〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 BHCS) MODE_IMM1 (Byte 〈x2,x9〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 BPL ) MODE_IMM1 (Byte 〈x2,xA〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 BMI ) MODE_IMM1 (Byte 〈x2,xB〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 BMC ) MODE_IMM1 (Byte 〈x2,xC〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 BMS ) MODE_IMM1 (Byte 〈x2,xD〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 BIL ) MODE_IMM1 (Byte 〈x2,xE〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 BIH ) MODE_IMM1 (Byte 〈x2,xF〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 BGE ) MODE_IMM1 (Byte 〈x9,x0〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 BLT ) MODE_IMM1 (Byte 〈x9,x1〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 BGT ) MODE_IMM1 (Byte 〈x9,x2〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 BLE ) MODE_IMM1 (Byte 〈x9,x3〉) 〈x0,x3〉
+].
+
+definition opcode_table_HC08_7 ≝
+[
+  quadrupleT ???? (anyOP HC08 BSETn) (MODE_DIRn o0) (Byte 〈x1,x0〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 BCLRn) (MODE_DIRn o0) (Byte 〈x1,x1〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 BSETn) (MODE_DIRn o1) (Byte 〈x1,x2〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 BCLRn) (MODE_DIRn o1) (Byte 〈x1,x3〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 BSETn) (MODE_DIRn o2) (Byte 〈x1,x4〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 BCLRn) (MODE_DIRn o2) (Byte 〈x1,x5〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 BSETn) (MODE_DIRn o3) (Byte 〈x1,x6〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 BCLRn) (MODE_DIRn o3) (Byte 〈x1,x7〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 BSETn) (MODE_DIRn o4) (Byte 〈x1,x8〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 BCLRn) (MODE_DIRn o4) (Byte 〈x1,x9〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 BSETn) (MODE_DIRn o5) (Byte 〈x1,xA〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 BCLRn) (MODE_DIRn o5) (Byte 〈x1,xB〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 BSETn) (MODE_DIRn o6) (Byte 〈x1,xC〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 BCLRn) (MODE_DIRn o6) (Byte 〈x1,xD〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 BSETn) (MODE_DIRn o7) (Byte 〈x1,xE〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 BCLRn) (MODE_DIRn o7) (Byte 〈x1,xF〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC08_8 ≝
+[
+  quadrupleT ???? (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x0〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x1〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x2〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x3〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x4〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x5〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x6〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x7〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x8〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x9〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xA〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xB〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xC〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xD〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 BRSETn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xE〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 BRCLRn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xF〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC08_9 ≝
+[
+  quadrupleT ???? (anyOP HC08 BIT) MODE_IMM1 (Byte 〈xA,x5〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 BIT) MODE_DIR1 (Byte 〈xB,x5〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 BIT) MODE_DIR2 (Byte 〈xC,x5〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 BIT) MODE_IX2  (Byte 〈xD,x5〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 BIT) MODE_IX1  (Byte 〈xE,x5〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 BIT) MODE_IX0  (Byte 〈xF,x5〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 BIT) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,x5〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 BIT) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,x5〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC08_10 ≝
+[
+  quadrupleT ???? (anyOP HC08 MUL ) MODE_INH  (Byte 〈x4,x2〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 DIV ) MODE_INH  (Byte 〈x5,x2〉) 〈x0,x7〉
+; quadrupleT ???? (anyOP HC08 NSA ) MODE_INH  (Byte 〈x6,x2〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 DAA ) MODE_INH  (Byte 〈x7,x2〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 RTI ) MODE_INH  (Byte 〈x8,x0〉) 〈x0,x7〉
+; quadrupleT ???? (anyOP HC08 RTS ) MODE_INH  (Byte 〈x8,x1〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 SWI ) MODE_INH  (Byte 〈x8,x3〉) 〈x0,x9〉
+; quadrupleT ???? (anyOP HC08 TAP ) MODE_INH  (Byte 〈x8,x4〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 TPA ) MODE_INH  (Byte 〈x8,x5〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 PULA) MODE_INH  (Byte 〈x8,x6〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 PSHA) MODE_INH  (Byte 〈x8,x7〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 PULX) MODE_INH  (Byte 〈x8,x8〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 PSHX) MODE_INH  (Byte 〈x8,x9〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 PULH) MODE_INH  (Byte 〈x8,xA〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 PSHH) MODE_INH  (Byte 〈x8,xB〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 STOP) MODE_INH  (Byte 〈x8,xE〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 WAIT) MODE_INH  (Byte 〈x8,xF〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 TXS ) MODE_INH  (Byte 〈x9,x4〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 TSX ) MODE_INH  (Byte 〈x9,x5〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 TAX ) MODE_INH  (Byte 〈x9,x7〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 CLC ) MODE_INH  (Byte 〈x9,x8〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 SEC ) MODE_INH  (Byte 〈x9,x9〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 CLI ) MODE_INH  (Byte 〈x9,xA〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 SEI ) MODE_INH  (Byte 〈x9,xB〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 RSP ) MODE_INH  (Byte 〈x9,xC〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 NOP ) MODE_INH  (Byte 〈x9,xD〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 TXA ) MODE_INH  (Byte 〈x9,xF〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 AIS ) MODE_IMM1 (Byte 〈xA,x7〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 AIX ) MODE_IMM1 (Byte 〈xA,xF〉) 〈x0,x2〉
+].
+
+definition opcode_table_HC08_11 ≝
+[
+  quadrupleT ???? (anyOP HC08 CBEQA) MODE_DIR1_and_IMM1 (Byte 〈x3,x1〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 CBEQA) MODE_IMM1_and_IMM1 (Byte 〈x4,x1〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 CBEQX) MODE_IMM1_and_IMM1 (Byte 〈x5,x1〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 CBEQA) MODE_IX1p_and_IMM1 (Byte 〈x6,x1〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 CBEQA) MODE_IX0p_and_IMM1 (Byte 〈x7,x1〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 CBEQA) MODE_SP1_and_IMM1  (Word 〈〈x9,xE〉:〈x6,x1〉〉) 〈x0,x6〉
+].
+
+definition opcode_table_HC08_12 ≝
+[
+  quadrupleT ???? (anyOP HC08 CLR) MODE_DIR1 (Byte 〈x3,xF〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 CLR) MODE_INHA (Byte 〈x4,xF〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 CLR) MODE_INHX (Byte 〈x5,xF〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 CLR) MODE_IX1  (Byte 〈x6,xF〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 CLR) MODE_IX0  (Byte 〈x7,xF〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 CLR) MODE_INHH (Byte 〈x8,xC〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 CLR) MODE_SP1  (Word 〈〈x9,xE〉:〈x6,xF〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC08_13 ≝
+[
+  quadrupleT ???? (anyOP HC08 CMP) MODE_IMM1 (Byte 〈xA,x1〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 CMP) MODE_DIR1 (Byte 〈xB,x1〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 CMP) MODE_DIR2 (Byte 〈xC,x1〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 CMP) MODE_IX2  (Byte 〈xD,x1〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 CMP) MODE_IX1  (Byte 〈xE,x1〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 CMP) MODE_IX0  (Byte 〈xF,x1〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 CMP) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,x1〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 CMP) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,x1〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC08_14 ≝
+[
+  quadrupleT ???? (anyOP HC08 COM) MODE_DIR1 (Byte 〈x3,x3〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 COM) MODE_INHA (Byte 〈x4,x3〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 COM) MODE_INHX (Byte 〈x5,x3〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 COM) MODE_IX1  (Byte 〈x6,x3〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 COM) MODE_IX0  (Byte 〈x7,x3〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 COM) MODE_SP1  (Word 〈〈x9,xE〉:〈x6,x3〉〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC08_15 ≝
+[
+  quadrupleT ???? (anyOP HC08 STHX) MODE_DIR1 (Byte 〈x3,x5〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 LDHX) MODE_IMM2 (Byte 〈x4,x5〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 LDHX) MODE_DIR1 (Byte 〈x5,x5〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 CPHX) MODE_IMM2 (Byte 〈x6,x5〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 CPHX) MODE_DIR1 (Byte 〈x7,x5〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC08_16 ≝
+[
+  quadrupleT ???? (anyOP HC08 CPX) MODE_IMM1 (Byte 〈xA,x3〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 CPX) MODE_DIR1 (Byte 〈xB,x3〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 CPX) MODE_DIR2 (Byte 〈xC,x3〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 CPX) MODE_IX2  (Byte 〈xD,x3〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 CPX) MODE_IX1  (Byte 〈xE,x3〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 CPX) MODE_IX0  (Byte 〈xF,x3〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 CPX) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,x3〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 CPX) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,x3〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC08_17 ≝
+[
+  quadrupleT ???? (anyOP HC08 DBNZ) MODE_DIR1_and_IMM1 (Byte 〈x3,xB〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 DBNZ) MODE_INHA_and_IMM1 (Byte 〈x4,xB〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 DBNZ) MODE_INHX_and_IMM1 (Byte 〈x5,xB〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 DBNZ) MODE_IX1_and_IMM1  (Byte 〈x6,xB〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 DBNZ) MODE_IX0_and_IMM1  (Byte 〈x7,xB〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 DBNZ) MODE_SP1_and_IMM1  (Word 〈〈x9,xE〉:〈x6,xB〉〉) 〈x0,x6〉
+].
+
+definition opcode_table_HC08_18 ≝
+[
+  quadrupleT ???? (anyOP HC08 DEC) MODE_DIR1 (Byte 〈x3,xA〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 DEC) MODE_INHA (Byte 〈x4,xA〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 DEC) MODE_INHX (Byte 〈x5,xA〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 DEC) MODE_IX1  (Byte 〈x6,xA〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 DEC) MODE_IX0  (Byte 〈x7,xA〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 DEC) MODE_SP1  (Word 〈〈x9,xE〉:〈x6,xA〉〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC08_19 ≝
+[
+  quadrupleT ???? (anyOP HC08 EOR) MODE_IMM1 (Byte 〈xA,x8〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 EOR) MODE_DIR1 (Byte 〈xB,x8〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 EOR) MODE_DIR2 (Byte 〈xC,x8〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 EOR) MODE_IX2  (Byte 〈xD,x8〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 EOR) MODE_IX1  (Byte 〈xE,x8〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 EOR) MODE_IX0  (Byte 〈xF,x8〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 EOR) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,x8〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 EOR) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,x8〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC08_20 ≝
+[
+  quadrupleT ???? (anyOP HC08 INC) MODE_DIR1 (Byte 〈x3,xC〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 INC) MODE_INHA (Byte 〈x4,xC〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 INC) MODE_INHX (Byte 〈x5,xC〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 INC) MODE_IX1  (Byte 〈x6,xC〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 INC) MODE_IX0  (Byte 〈x7,xC〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 INC) MODE_SP1  (Word 〈〈x9,xE〉:〈x6,xC〉〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC08_21 ≝
+[
+  quadrupleT ???? (anyOP HC08 JMP) MODE_DIR1 (Byte 〈xB,xC〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 JMP) MODE_DIR2 (Byte 〈xC,xC〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 JMP) MODE_IX2  (Byte 〈xD,xC〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 JMP) MODE_IX1  (Byte 〈xE,xC〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 JMP) MODE_IX0  (Byte 〈xF,xC〉) 〈x0,x3〉
+].
+
+definition opcode_table_HC08_22 ≝
+[
+  quadrupleT ???? (anyOP HC08 BSR) MODE_IMM1 (Byte 〈xA,xD〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 JSR) MODE_DIR1 (Byte 〈xB,xD〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 JSR) MODE_DIR2 (Byte 〈xC,xD〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 JSR) MODE_IX2  (Byte 〈xD,xD〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HC08 JSR) MODE_IX1  (Byte 〈xE,xD〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 JSR) MODE_IX0  (Byte 〈xF,xD〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC08_23 ≝
+[
+  quadrupleT ???? (anyOP HC08 LDA) MODE_IMM1 (Byte 〈xA,x6〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 LDA) MODE_DIR1 (Byte 〈xB,x6〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 LDA) MODE_DIR2 (Byte 〈xC,x6〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 LDA) MODE_IX2  (Byte 〈xD,x6〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 LDA) MODE_IX1  (Byte 〈xE,x6〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 LDA) MODE_IX0  (Byte 〈xF,x6〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 LDA) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,x6〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 LDA) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,x6〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC08_24 ≝
+[
+  quadrupleT ???? (anyOP HC08 LDX) MODE_IMM1 (Byte 〈xA,xE〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 LDX) MODE_DIR1 (Byte 〈xB,xE〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 LDX) MODE_DIR2 (Byte 〈xC,xE〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 LDX) MODE_IX2  (Byte 〈xD,xE〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 LDX) MODE_IX1  (Byte 〈xE,xE〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 LDX) MODE_IX0  (Byte 〈xF,xE〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 LDX) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,xE〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 LDX) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,xE〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC08_25 ≝
+[
+  quadrupleT ???? (anyOP HC08 LSR) MODE_DIR1 (Byte 〈x3,x4〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 LSR) MODE_INHA (Byte 〈x4,x4〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 LSR) MODE_INHX (Byte 〈x5,x4〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 LSR) MODE_IX1  (Byte 〈x6,x4〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 LSR) MODE_IX0  (Byte 〈x7,x4〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 LSR) MODE_SP1  (Word 〈〈x9,xE〉:〈x6,x4〉〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC08_26 ≝
+[
+  quadrupleT ???? (anyOP HC08 MOV) MODE_DIR1_to_DIR1 (Byte 〈x4,xE〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 MOV) MODE_DIR1_to_IX0p (Byte 〈x5,xE〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 MOV) MODE_IMM1_to_DIR1 (Byte 〈x6,xE〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 MOV) MODE_IX0p_to_DIR1 (Byte 〈x7,xE〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC08_27 ≝
+[
+  quadrupleT ???? (anyOP HC08 NEG) MODE_DIR1 (Byte 〈x3,x0〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 NEG) MODE_INHA (Byte 〈x4,x0〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 NEG) MODE_INHX (Byte 〈x5,x0〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 NEG) MODE_IX1  (Byte 〈x6,x0〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 NEG) MODE_IX0  (Byte 〈x7,x0〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 NEG) MODE_SP1  (Word 〈〈x9,xE〉:〈x6,x0〉〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC08_28 ≝
+[
+  quadrupleT ???? (anyOP HC08 ORA) MODE_IMM1 (Byte 〈xA,xA〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 ORA) MODE_DIR1 (Byte 〈xB,xA〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 ORA) MODE_DIR2 (Byte 〈xC,xA〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 ORA) MODE_IX2  (Byte 〈xD,xA〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 ORA) MODE_IX1  (Byte 〈xE,xA〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 ORA) MODE_IX0  (Byte 〈xF,xA〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 ORA) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,xA〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 ORA) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,xA〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC08_29 ≝
+[
+  quadrupleT ???? (anyOP HC08 ROL) MODE_DIR1 (Byte 〈x3,x9〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 ROL) MODE_INHA (Byte 〈x4,x9〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 ROL) MODE_INHX (Byte 〈x5,x9〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 ROL) MODE_IX1  (Byte 〈x6,x9〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 ROL) MODE_IX0  (Byte 〈x7,x9〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 ROL) MODE_SP1  (Word 〈〈x9,xE〉:〈x6,x9〉〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC08_30 ≝
+[
+  quadrupleT ???? (anyOP HC08 ROR) MODE_DIR1 (Byte 〈x3,x6〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 ROR) MODE_INHA (Byte 〈x4,x6〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 ROR) MODE_INHX (Byte 〈x5,x6〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 ROR) MODE_IX1  (Byte 〈x6,x6〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 ROR) MODE_IX0  (Byte 〈x7,x6〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 ROR) MODE_SP1  (Word 〈〈x9,xE〉:〈x6,x6〉〉) 〈x0,x5〉
+].
+
+definition opcode_table_HC08_31 ≝
+[
+  quadrupleT ???? (anyOP HC08 SBC) MODE_IMM1 (Byte 〈xA,x2〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 SBC) MODE_DIR1 (Byte 〈xB,x2〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 SBC) MODE_DIR2 (Byte 〈xC,x2〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 SBC) MODE_IX2  (Byte 〈xD,x2〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 SBC) MODE_IX1  (Byte 〈xE,x2〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 SBC) MODE_IX0  (Byte 〈xF,x2〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 SBC) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,x2〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 SBC) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,x2〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC08_32 ≝
+[
+  quadrupleT ???? (anyOP HC08 STA) MODE_DIR1 (Byte 〈xB,x7〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 STA) MODE_DIR2 (Byte 〈xC,x7〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 STA) MODE_IX2  (Byte 〈xD,x7〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 STA) MODE_IX1  (Byte 〈xE,x7〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 STA) MODE_IX0  (Byte 〈xF,x7〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 STA) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,x7〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 STA) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,x7〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC08_33 ≝
+[
+  quadrupleT ???? (anyOP HC08 STX) MODE_DIR1 (Byte 〈xB,xF〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 STX) MODE_DIR2 (Byte 〈xC,xF〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 STX) MODE_IX2  (Byte 〈xD,xF〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 STX) MODE_IX1  (Byte 〈xE,xF〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 STX) MODE_IX0  (Byte 〈xF,xF〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 STX) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,xF〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 STX) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,xF〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC08_34 ≝
+[
+  quadrupleT ???? (anyOP HC08 SUB) MODE_IMM1 (Byte 〈xA,x0〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 SUB) MODE_DIR1 (Byte 〈xB,x0〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 SUB) MODE_DIR2 (Byte 〈xC,x0〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 SUB) MODE_IX2  (Byte 〈xD,x0〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HC08 SUB) MODE_IX1  (Byte 〈xE,x0〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 SUB) MODE_IX0  (Byte 〈xF,x0〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 SUB) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,x0〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HC08 SUB) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,x0〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC08_35 ≝
+[
+  quadrupleT ???? (anyOP HC08 TST) MODE_DIR1 (Byte 〈x3,xD〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 TST) MODE_INHA (Byte 〈x4,xD〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 TST) MODE_INHX (Byte 〈x5,xD〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HC08 TST) MODE_IX1  (Byte 〈x6,xD〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HC08 TST) MODE_IX0  (Byte 〈x7,xD〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HC08 TST) MODE_SP1  (Word 〈〈x9,xE〉:〈x6,xD〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HC08 ≝
+opcode_table_HC08_1  @ opcode_table_HC08_2  @ opcode_table_HC08_3  @ opcode_table_HC08_4  @
+opcode_table_HC08_5  @ opcode_table_HC08_6  @ opcode_table_HC08_7  @ opcode_table_HC08_8  @
+opcode_table_HC08_9  @ opcode_table_HC08_10 @ opcode_table_HC08_11 @ opcode_table_HC08_12 @
+opcode_table_HC08_13 @ opcode_table_HC08_14 @ opcode_table_HC08_15 @ opcode_table_HC08_16 @
+opcode_table_HC08_17 @ opcode_table_HC08_18 @ opcode_table_HC08_19 @ opcode_table_HC08_20 @
+opcode_table_HC08_21 @ opcode_table_HC08_22 @ opcode_table_HC08_23 @ opcode_table_HC08_24 @
+opcode_table_HC08_25 @ opcode_table_HC08_26 @ opcode_table_HC08_27 @ opcode_table_HC08_28 @
+opcode_table_HC08_29 @ opcode_table_HC08_30 @ opcode_table_HC08_31 @ opcode_table_HC08_32 @
+opcode_table_HC08_33 @ opcode_table_HC08_34 @ opcode_table_HC08_35.
+
+(* CORRETTEZZA *)
+
+(* HC08: opcode non implementati come da manuale (byte) *)
+definition HC08_not_impl_1 ≝
+ [〈x3,x2〉;〈x3,xE〉
+ ;〈x8,x2〉;〈x8,xD〉
+ ;〈x9,x6〉;〈x9,xE〉
+ ;〈xA,xC〉
+ ].
+
+lemma ok_byte_table_HC08 : forall_byte8 (λb.
+ (test_not_impl b HC08_not_impl_1     ⊕ eqb (get_byte_count HC08 b 0 opcode_table_HC08) 1) ⊗
+ (⊖ (test_not_impl b HC08_not_impl_1) ⊕ eqb (get_byte_count HC08 b 0 opcode_table_HC08) 0))
+ = true.
+ reflexivity.
+qed.
+
+(* HC08: opcode non implementati come da manuale (0x9E+byte) *)
+definition HC08_not_impl_2 ≝
+ [〈x0,x0〉;〈x0,x1〉;〈x0,x2〉;〈x0,x3〉;〈x0,x4〉;〈x0,x5〉;〈x0,x6〉;〈x0,x7〉
+ ;〈x0,x8〉;〈x0,x9〉;〈x0,xA〉;〈x0,xB〉;〈x0,xC〉;〈x0,xD〉;〈x0,xE〉;〈x0,xF〉
+ ;〈x1,x0〉;〈x1,x1〉;〈x1,x2〉;〈x1,x3〉;〈x1,x4〉;〈x1,x5〉;〈x1,x6〉;〈x1,x7〉
+ ;〈x1,x8〉;〈x1,x9〉;〈x1,xA〉;〈x1,xB〉;〈x1,xC〉;〈x1,xD〉;〈x1,xE〉;〈x1,xF〉
+ ;〈x2,x0〉;〈x2,x1〉;〈x2,x2〉;〈x2,x3〉;〈x2,x4〉;〈x2,x5〉;〈x2,x6〉;〈x2,x7〉
+ ;〈x2,x8〉;〈x2,x9〉;〈x2,xA〉;〈x2,xB〉;〈x2,xC〉;〈x2,xD〉;〈x2,xE〉;〈x2,xF〉
+ ;〈x3,x0〉;〈x3,x1〉;〈x3,x2〉;〈x3,x3〉;〈x3,x4〉;〈x3,x5〉;〈x3,x6〉;〈x3,x7〉
+ ;〈x3,x8〉;〈x3,x9〉;〈x3,xA〉;〈x3,xB〉;〈x3,xC〉;〈x3,xD〉;〈x3,xE〉;〈x3,xF〉
+ ;〈x4,x0〉;〈x4,x1〉;〈x4,x2〉;〈x4,x3〉;〈x4,x4〉;〈x4,x5〉;〈x4,x6〉;〈x4,x7〉
+ ;〈x4,x8〉;〈x4,x9〉;〈x4,xA〉;〈x4,xB〉;〈x4,xC〉;〈x4,xD〉;〈x4,xE〉;〈x4,xF〉
+ ;〈x5,x0〉;〈x5,x1〉;〈x5,x2〉;〈x5,x3〉;〈x5,x4〉;〈x5,x5〉;〈x5,x6〉;〈x5,x7〉
+ ;〈x5,x8〉;〈x5,x9〉;〈x5,xA〉;〈x5,xB〉;〈x5,xC〉;〈x5,xD〉;〈x5,xE〉;〈x5,xF〉
+ ;〈x6,x2〉;〈x6,x5〉;〈x6,xE〉
+ ;〈x7,x0〉;〈x7,x1〉;〈x7,x2〉;〈x7,x3〉;〈x7,x4〉;〈x7,x5〉;〈x7,x6〉;〈x7,x7〉
+ ;〈x7,x8〉;〈x7,x9〉;〈x7,xA〉;〈x7,xB〉;〈x7,xC〉;〈x7,xD〉;〈x7,xE〉;〈x7,xF〉
+ ;〈x8,x0〉;〈x8,x1〉;〈x8,x2〉;〈x8,x3〉;〈x8,x4〉;〈x8,x5〉;〈x8,x6〉;〈x8,x7〉
+ ;〈x8,x8〉;〈x8,x9〉;〈x8,xA〉;〈x8,xB〉;〈x8,xC〉;〈x8,xD〉;〈x8,xE〉;〈x8,xF〉
+ ;〈x9,x0〉;〈x9,x1〉;〈x9,x2〉;〈x9,x3〉;〈x9,x4〉;〈x9,x5〉;〈x9,x6〉;〈x9,x7〉
+ ;〈x9,x8〉;〈x9,x9〉;〈x9,xA〉;〈x9,xB〉;〈x9,xC〉;〈x9,xD〉;〈x9,xE〉;〈x9,xF〉
+ ;〈xA,x0〉;〈xA,x1〉;〈xA,x2〉;〈xA,x3〉;〈xA,x4〉;〈xA,x5〉;〈xA,x6〉;〈xA,x7〉
+ ;〈xA,x8〉;〈xA,x9〉;〈xA,xA〉;〈xA,xB〉;〈xA,xC〉;〈xA,xD〉;〈xA,xE〉;〈xA,xF〉
+ ;〈xB,x0〉;〈xB,x1〉;〈xB,x2〉;〈xB,x3〉;〈xB,x4〉;〈xB,x5〉;〈xB,x6〉;〈xB,x7〉
+ ;〈xB,x8〉;〈xB,x9〉;〈xB,xA〉;〈xB,xB〉;〈xB,xC〉;〈xB,xD〉;〈xB,xE〉;〈xB,xF〉
+ ;〈xC,x0〉;〈xC,x1〉;〈xC,x2〉;〈xC,x3〉;〈xC,x4〉;〈xC,x5〉;〈xC,x6〉;〈xC,x7〉
+ ;〈xC,x8〉;〈xC,x9〉;〈xC,xA〉;〈xC,xB〉;〈xC,xC〉;〈xC,xD〉;〈xC,xE〉;〈xC,xF〉
+ ;〈xD,xC〉;〈xD,xD〉
+ ;〈xE,xC〉;〈xE,xD〉
+ ;〈xF,x0〉;〈xF,x1〉;〈xF,x2〉;〈xF,x3〉;〈xF,x4〉;〈xF,x5〉;〈xF,x6〉;〈xF,x7〉
+ ;〈xF,x8〉;〈xF,x9〉;〈xF,xA〉;〈xF,xB〉;〈xF,xC〉;〈xF,xD〉;〈xF,xE〉;〈xF,xF〉
+ ].
+
+lemma ok_byte_table_HC08_2 : forall_byte8 (λb.
+ (test_not_impl b HC08_not_impl_2     ⊕ eqb (get_word_count HC08 b 0 opcode_table_HC08) 1) ⊗
+ (⊖ (test_not_impl b HC08_not_impl_2) ⊕ eqb (get_word_count HC08 b 0 opcode_table_HC08) 0))
+ = true.
+ reflexivity.
+qed.
+
+(*
+lemma ok_OpIm_table_HC08 :
+ forall_instr_mode (λi:instr_mode.
+ forall_opcode     (λop:opcode.
+  leb (get_OpIm_count HC08 (anyOP HC08 op) i 0 opcode_table_HC08) 1)) = true.
+ reflexivity.
+qed.
+*)
diff --git a/helm/software/matita/library/freescale/table_HCS08.ma b/helm/software/matita/library/freescale/table_HCS08.ma
new file mode 100644 (file)
index 0000000..bcc8525
--- /dev/null
@@ -0,0 +1,553 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/table_HCS08/".
+
+(*include "/media/VIRTUOSO/freescale/opcode.ma".*)
+include "freescale/opcode.ma".
+
+(* ****************** *)
+(* TABELLA DELL'HCS08 *)
+(* ****************** *)
+
+(* definizione come concatenazione finale di liste per velocizzare il parsing *)
+(* ogni riga e' (any_opcode m) (instr_mode) (opcode esadecimale) (#cicli esecuzione) *)
+(* NB: l'uso di any_opcode m + concatenazione finale tutte liste
+       impedisce di introdurre opcode disomogenei (per mcu) *)
+
+definition opcode_table_HCS08_1 ≝
+[
+  quadrupleT ???? (anyOP HCS08 ADC) MODE_IMM1 (Byte 〈xA,x9〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HCS08 ADC) MODE_DIR1 (Byte 〈xB,x9〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 ADC) MODE_DIR2 (Byte 〈xC,x9〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 ADC) MODE_IX2  (Byte 〈xD,x9〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 ADC) MODE_IX1  (Byte 〈xE,x9〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 ADC) MODE_IX0  (Byte 〈xF,x9〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 ADC) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,x9〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 ADC) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,x9〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HCS08_2 ≝
+[
+  quadrupleT ???? (anyOP HCS08 ADD) MODE_IMM1 (Byte 〈xA,xB〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HCS08 ADD) MODE_DIR1 (Byte 〈xB,xB〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 ADD) MODE_DIR2 (Byte 〈xC,xB〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 ADD) MODE_IX2  (Byte 〈xD,xB〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 ADD) MODE_IX1  (Byte 〈xE,xB〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 ADD) MODE_IX0  (Byte 〈xF,xB〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 ADD) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,xB〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 ADD) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,xB〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HCS08_3 ≝
+[
+  quadrupleT ???? (anyOP HCS08 AND) MODE_IMM1 (Byte 〈xA,x4〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HCS08 AND) MODE_DIR1 (Byte 〈xB,x4〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 AND) MODE_DIR2 (Byte 〈xC,x4〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 AND) MODE_IX2  (Byte 〈xD,x4〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 AND) MODE_IX1  (Byte 〈xE,x4〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 AND) MODE_IX0  (Byte 〈xF,x4〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 AND) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,x4〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 AND) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,x4〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HCS08_4 ≝
+[
+  quadrupleT ???? (anyOP HCS08 ASL) MODE_DIR1 (Byte 〈x3,x8〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 ASL) MODE_INHA (Byte 〈x4,x8〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 ASL) MODE_INHX (Byte 〈x5,x8〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 ASL) MODE_IX1  (Byte 〈x6,x8〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 ASL) MODE_IX0  (Byte 〈x7,x8〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 ASL) MODE_SP1  (Word 〈〈x9,xE〉:〈x6,x8〉〉) 〈x0,x6〉
+].
+
+definition opcode_table_HCS08_5 ≝
+[
+  quadrupleT ???? (anyOP HCS08 ASR) MODE_DIR1 (Byte 〈x3,x7〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 ASR) MODE_INHA (Byte 〈x4,x7〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 ASR) MODE_INHX (Byte 〈x5,x7〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 ASR) MODE_IX1  (Byte 〈x6,x7〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 ASR) MODE_IX0  (Byte 〈x7,x7〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 ASR) MODE_SP1  (Word 〈〈x9,xE〉:〈x6,x7〉〉) 〈x0,x6〉
+].
+
+definition opcode_table_HCS08_6 ≝
+[
+  quadrupleT ???? (anyOP HCS08 BRA ) MODE_IMM1 (Byte 〈x2,x0〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 BRN ) MODE_IMM1 (Byte 〈x2,x1〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 BHI ) MODE_IMM1 (Byte 〈x2,x2〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 BLS ) MODE_IMM1 (Byte 〈x2,x3〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 BCC ) MODE_IMM1 (Byte 〈x2,x4〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 BCS ) MODE_IMM1 (Byte 〈x2,x5〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 BNE ) MODE_IMM1 (Byte 〈x2,x6〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 BEQ ) MODE_IMM1 (Byte 〈x2,x7〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 BHCC) MODE_IMM1 (Byte 〈x2,x8〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 BHCS) MODE_IMM1 (Byte 〈x2,x9〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 BPL ) MODE_IMM1 (Byte 〈x2,xA〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 BMI ) MODE_IMM1 (Byte 〈x2,xB〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 BMC ) MODE_IMM1 (Byte 〈x2,xC〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 BMS ) MODE_IMM1 (Byte 〈x2,xD〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 BIL ) MODE_IMM1 (Byte 〈x2,xE〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 BIH ) MODE_IMM1 (Byte 〈x2,xF〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 BGE ) MODE_IMM1 (Byte 〈x9,x0〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 BLT ) MODE_IMM1 (Byte 〈x9,x1〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 BGT ) MODE_IMM1 (Byte 〈x9,x2〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 BLE ) MODE_IMM1 (Byte 〈x9,x3〉) 〈x0,x3〉
+].
+
+definition opcode_table_HCS08_7 ≝
+[
+  quadrupleT ???? (anyOP HCS08 BSETn) (MODE_DIRn o0) (Byte 〈x1,x0〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BCLRn) (MODE_DIRn o0) (Byte 〈x1,x1〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BSETn) (MODE_DIRn o1) (Byte 〈x1,x2〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BCLRn) (MODE_DIRn o1) (Byte 〈x1,x3〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BSETn) (MODE_DIRn o2) (Byte 〈x1,x4〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BCLRn) (MODE_DIRn o2) (Byte 〈x1,x5〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BSETn) (MODE_DIRn o3) (Byte 〈x1,x6〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BCLRn) (MODE_DIRn o3) (Byte 〈x1,x7〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BSETn) (MODE_DIRn o4) (Byte 〈x1,x8〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BCLRn) (MODE_DIRn o4) (Byte 〈x1,x9〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BSETn) (MODE_DIRn o5) (Byte 〈x1,xA〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BCLRn) (MODE_DIRn o5) (Byte 〈x1,xB〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BSETn) (MODE_DIRn o6) (Byte 〈x1,xC〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BCLRn) (MODE_DIRn o6) (Byte 〈x1,xD〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BSETn) (MODE_DIRn o7) (Byte 〈x1,xE〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BCLRn) (MODE_DIRn o7) (Byte 〈x1,xF〉) 〈x0,x5〉
+].
+
+definition opcode_table_HCS08_8 ≝
+[
+  quadrupleT ???? (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x0〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x1〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x2〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x3〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x4〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x5〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x6〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x7〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x8〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x9〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xA〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xB〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xC〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xD〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BRSETn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xE〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BRCLRn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xF〉) 〈x0,x5〉
+].
+
+definition opcode_table_HCS08_9 ≝
+[
+  quadrupleT ???? (anyOP HCS08 BIT) MODE_IMM1 (Byte 〈xA,x5〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HCS08 BIT) MODE_DIR1 (Byte 〈xB,x5〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 BIT) MODE_DIR2 (Byte 〈xC,x5〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 BIT) MODE_IX2  (Byte 〈xD,x5〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 BIT) MODE_IX1  (Byte 〈xE,x5〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 BIT) MODE_IX0  (Byte 〈xF,x5〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 BIT) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,x5〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 BIT) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,x5〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HCS08_10 ≝
+[
+  quadrupleT ???? (anyOP HCS08 MUL ) MODE_INH  (Byte 〈x4,x2〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 DIV ) MODE_INH  (Byte 〈x5,x2〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HCS08 NSA ) MODE_INH  (Byte 〈x6,x2〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 DAA ) MODE_INH  (Byte 〈x7,x2〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 RTI ) MODE_INH  (Byte 〈x8,x0〉) 〈x0,x9〉
+; quadrupleT ???? (anyOP HCS08 RTS ) MODE_INH  (Byte 〈x8,x1〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HCS08 SWI ) MODE_INH  (Byte 〈x8,x3〉) 〈x0,xB〉
+; quadrupleT ???? (anyOP HCS08 BGND) MODE_INH  (Byte 〈x8,x2〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 TAP ) MODE_INH  (Byte 〈x8,x4〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 TPA ) MODE_INH  (Byte 〈x8,x5〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 PULA) MODE_INH  (Byte 〈x8,x6〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 PSHA) MODE_INH  (Byte 〈x8,x7〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HCS08 PULX) MODE_INH  (Byte 〈x8,x8〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 PSHX) MODE_INH  (Byte 〈x8,x9〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HCS08 PULH) MODE_INH  (Byte 〈x8,xA〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 PSHH) MODE_INH  (Byte 〈x8,xB〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HCS08 STOP) MODE_INH  (Byte 〈x8,xE〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HCS08 WAIT) MODE_INH  (Byte 〈x8,xF〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HCS08 TXS ) MODE_INH  (Byte 〈x9,x4〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HCS08 TSX ) MODE_INH  (Byte 〈x9,x5〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HCS08 TAX ) MODE_INH  (Byte 〈x9,x7〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 CLC ) MODE_INH  (Byte 〈x9,x8〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 SEC ) MODE_INH  (Byte 〈x9,x9〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 CLI ) MODE_INH  (Byte 〈x9,xA〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 SEI ) MODE_INH  (Byte 〈x9,xB〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 RSP ) MODE_INH  (Byte 〈x9,xC〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 NOP ) MODE_INH  (Byte 〈x9,xD〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 TXA ) MODE_INH  (Byte 〈x9,xF〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 AIS ) MODE_IMM1 (Byte 〈xA,x7〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HCS08 AIX ) MODE_IMM1 (Byte 〈xA,xF〉) 〈x0,x2〉
+].
+
+definition opcode_table_HCS08_11 ≝
+[
+  quadrupleT ???? (anyOP HCS08 CBEQA) MODE_DIR1_and_IMM1 (Byte 〈x3,x1〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 CBEQA) MODE_IMM1_and_IMM1 (Byte 〈x4,x1〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 CBEQX) MODE_IMM1_and_IMM1 (Byte 〈x5,x1〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 CBEQA) MODE_IX1p_and_IMM1 (Byte 〈x6,x1〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 CBEQA) MODE_IX0p_and_IMM1 (Byte 〈x7,x1〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 CBEQA) MODE_SP1_and_IMM1  (Word 〈〈x9,xE〉:〈x6,x1〉〉) 〈x0,x6〉
+].
+
+definition opcode_table_HCS08_12 ≝
+[
+  quadrupleT ???? (anyOP HCS08 CLR) MODE_DIR1 (Byte 〈x3,xF〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 CLR) MODE_INHA (Byte 〈x4,xF〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 CLR) MODE_INHX (Byte 〈x5,xF〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 CLR) MODE_IX1  (Byte 〈x6,xF〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 CLR) MODE_IX0  (Byte 〈x7,xF〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 CLR) MODE_INHH (Byte 〈x8,xC〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 CLR) MODE_SP1  (Word 〈〈x9,xE〉:〈x6,xF〉〉) 〈x0,x6〉
+].
+
+definition opcode_table_HCS08_13 ≝
+[
+  quadrupleT ???? (anyOP HCS08 CMP) MODE_IMM1 (Byte 〈xA,x1〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HCS08 CMP) MODE_DIR1 (Byte 〈xB,x1〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 CMP) MODE_DIR2 (Byte 〈xC,x1〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 CMP) MODE_IX2  (Byte 〈xD,x1〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 CMP) MODE_IX1  (Byte 〈xE,x1〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 CMP) MODE_IX0  (Byte 〈xF,x1〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 CMP) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,x1〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 CMP) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,x1〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HCS08_14 ≝
+[
+  quadrupleT ???? (anyOP HCS08 COM) MODE_DIR1 (Byte 〈x3,x3〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 COM) MODE_INHA (Byte 〈x4,x3〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 COM) MODE_INHX (Byte 〈x5,x3〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 COM) MODE_IX1  (Byte 〈x6,x3〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 COM) MODE_IX0  (Byte 〈x7,x3〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 COM) MODE_SP1  (Word 〈〈x9,xE〉:〈x6,x3〉〉) 〈x0,x6〉
+].
+
+definition opcode_table_HCS08_15 ≝
+[
+  quadrupleT ???? (anyOP HCS08 CPHX) MODE_DIR2 (Byte 〈x3,xE〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HCS08 CPHX) MODE_IMM2 (Byte 〈x6,x5〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 CPHX) MODE_DIR1 (Byte 〈x7,x5〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 CPHX) MODE_SP1  (Word 〈〈x9,xE〉:〈xF,x3〉〉) 〈x0,x6〉
+
+; quadrupleT ???? (anyOP HCS08 LDHX) MODE_DIR2 (Byte 〈x3,x2〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 LDHX) MODE_IMM2 (Byte 〈x4,x5〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 LDHX) MODE_DIR1 (Byte 〈x5,x5〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 LDHX) MODE_IX0  (Word 〈〈x9,xE〉:〈xA,xE〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 LDHX) MODE_IX2  (Word 〈〈x9,xE〉:〈xB,xE〉〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HCS08 LDHX) MODE_IX1  (Word 〈〈x9,xE〉:〈xC,xE〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 LDHX) MODE_SP1  (Word 〈〈x9,xE〉:〈xF,xE〉〉) 〈x0,x5〉
+
+; quadrupleT ???? (anyOP HCS08 STHX) MODE_DIR1 (Byte 〈x3,x5〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 STHX) MODE_DIR2 (Byte 〈x9,x6〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 STHX) MODE_SP1  (Word 〈〈x9,xE〉:〈xF,xF〉〉) 〈x0,x5〉
+].
+
+definition opcode_table_HCS08_16 ≝
+[
+  quadrupleT ???? (anyOP HCS08 CPX) MODE_IMM1 (Byte 〈xA,x3〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HCS08 CPX) MODE_DIR1 (Byte 〈xB,x3〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 CPX) MODE_DIR2 (Byte 〈xC,x3〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 CPX) MODE_IX2  (Byte 〈xD,x3〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 CPX) MODE_IX1  (Byte 〈xE,x3〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 CPX) MODE_IX0  (Byte 〈xF,x3〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 CPX) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,x3〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 CPX) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,x3〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HCS08_17 ≝
+[
+  quadrupleT ???? (anyOP HCS08 DBNZ) MODE_DIR1_and_IMM1 (Byte 〈x3,xB〉) 〈x0,x7〉
+; quadrupleT ???? (anyOP HCS08 DBNZ) MODE_INHA_and_IMM1 (Byte 〈x4,xB〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 DBNZ) MODE_INHX_and_IMM1 (Byte 〈x5,xB〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 DBNZ) MODE_IX1_and_IMM1  (Byte 〈x6,xB〉) 〈x0,x7〉
+; quadrupleT ???? (anyOP HCS08 DBNZ) MODE_IX0_and_IMM1  (Byte 〈x7,xB〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HCS08 DBNZ) MODE_SP1_and_IMM1  (Word 〈〈x9,xE〉:〈x6,xB〉〉) 〈x0,x8〉
+].
+
+definition opcode_table_HCS08_18 ≝
+[
+  quadrupleT ???? (anyOP HCS08 DEC) MODE_DIR1 (Byte 〈x3,xA〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 DEC) MODE_INHA (Byte 〈x4,xA〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 DEC) MODE_INHX (Byte 〈x5,xA〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 DEC) MODE_IX1  (Byte 〈x6,xA〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 DEC) MODE_IX0  (Byte 〈x7,xA〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 DEC) MODE_SP1  (Word 〈〈x9,xE〉:〈x6,xA〉〉) 〈x0,x6〉
+].
+
+definition opcode_table_HCS08_19 ≝
+[
+  quadrupleT ???? (anyOP HCS08 EOR) MODE_IMM1 (Byte 〈xA,x8〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HCS08 EOR) MODE_DIR1 (Byte 〈xB,x8〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 EOR) MODE_DIR2 (Byte 〈xC,x8〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 EOR) MODE_IX2  (Byte 〈xD,x8〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 EOR) MODE_IX1  (Byte 〈xE,x8〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 EOR) MODE_IX0  (Byte 〈xF,x8〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 EOR) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,x8〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 EOR) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,x8〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HCS08_20 ≝
+[
+  quadrupleT ???? (anyOP HCS08 INC) MODE_DIR1 (Byte 〈x3,xC〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 INC) MODE_INHA (Byte 〈x4,xC〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 INC) MODE_INHX (Byte 〈x5,xC〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 INC) MODE_IX1  (Byte 〈x6,xC〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 INC) MODE_IX0  (Byte 〈x7,xC〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 INC) MODE_SP1  (Word 〈〈x9,xE〉:〈x6,xC〉〉) 〈x0,x6〉
+].
+
+definition opcode_table_HCS08_21 ≝
+[
+  quadrupleT ???? (anyOP HCS08 JMP) MODE_DIR1 (Byte 〈xB,xC〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 JMP) MODE_DIR2 (Byte 〈xC,xC〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 JMP) MODE_IX2  (Byte 〈xD,xC〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 JMP) MODE_IX1  (Byte 〈xE,xC〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 JMP) MODE_IX0  (Byte 〈xF,xC〉) 〈x0,x3〉
+].
+
+definition opcode_table_HCS08_22 ≝
+[
+  quadrupleT ???? (anyOP HCS08 BSR) MODE_IMM1 (Byte 〈xA,xD〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 JSR) MODE_DIR1 (Byte 〈xB,xD〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 JSR) MODE_DIR2 (Byte 〈xC,xD〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HCS08 JSR) MODE_IX2  (Byte 〈xD,xD〉) 〈x0,x6〉
+; quadrupleT ???? (anyOP HCS08 JSR) MODE_IX1  (Byte 〈xE,xD〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 JSR) MODE_IX0  (Byte 〈xF,xD〉) 〈x0,x5〉
+].
+
+definition opcode_table_HCS08_23 ≝
+[
+  quadrupleT ???? (anyOP HCS08 LDA) MODE_IMM1 (Byte 〈xA,x6〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HCS08 LDA) MODE_DIR1 (Byte 〈xB,x6〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 LDA) MODE_DIR2 (Byte 〈xC,x6〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 LDA) MODE_IX2  (Byte 〈xD,x6〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 LDA) MODE_IX1  (Byte 〈xE,x6〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 LDA) MODE_IX0  (Byte 〈xF,x6〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 LDA) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,x6〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 LDA) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,x6〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HCS08_24 ≝
+[
+  quadrupleT ???? (anyOP HCS08 LDX) MODE_IMM1 (Byte 〈xA,xE〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HCS08 LDX) MODE_DIR1 (Byte 〈xB,xE〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 LDX) MODE_DIR2 (Byte 〈xC,xE〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 LDX) MODE_IX2  (Byte 〈xD,xE〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 LDX) MODE_IX1  (Byte 〈xE,xE〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 LDX) MODE_IX0  (Byte 〈xF,xE〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 LDX) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,xE〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 LDX) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,xE〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HCS08_25 ≝
+[
+  quadrupleT ???? (anyOP HCS08 LSR) MODE_DIR1 (Byte 〈x3,x4〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 LSR) MODE_INHA (Byte 〈x4,x4〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 LSR) MODE_INHX (Byte 〈x5,x4〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 LSR) MODE_IX1  (Byte 〈x6,x4〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 LSR) MODE_IX0  (Byte 〈x7,x4〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 LSR) MODE_SP1  (Word 〈〈x9,xE〉:〈x6,x4〉〉) 〈x0,x6〉
+].
+
+definition opcode_table_HCS08_26 ≝
+[
+  quadrupleT ???? (anyOP HCS08 MOV) MODE_DIR1_to_DIR1 (Byte 〈x4,xE〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 MOV) MODE_DIR1_to_IX0p (Byte 〈x5,xE〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 MOV) MODE_IMM1_to_DIR1 (Byte 〈x6,xE〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 MOV) MODE_IX0p_to_DIR1 (Byte 〈x7,xE〉) 〈x0,x5〉
+].
+
+definition opcode_table_HCS08_27 ≝
+[
+  quadrupleT ???? (anyOP HCS08 NEG) MODE_DIR1 (Byte 〈x3,x0〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 NEG) MODE_INHA (Byte 〈x4,x0〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 NEG) MODE_INHX (Byte 〈x5,x0〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 NEG) MODE_IX1  (Byte 〈x6,x0〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 NEG) MODE_IX0  (Byte 〈x7,x0〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 NEG) MODE_SP1  (Word 〈〈x9,xE〉:〈x6,x0〉〉) 〈x0,x6〉
+].
+
+definition opcode_table_HCS08_28 ≝
+[
+  quadrupleT ???? (anyOP HCS08 ORA) MODE_IMM1 (Byte 〈xA,xA〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HCS08 ORA) MODE_DIR1 (Byte 〈xB,xA〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 ORA) MODE_DIR2 (Byte 〈xC,xA〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 ORA) MODE_IX2  (Byte 〈xD,xA〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 ORA) MODE_IX1  (Byte 〈xE,xA〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 ORA) MODE_IX0  (Byte 〈xF,xA〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 ORA) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,xA〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 ORA) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,xA〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HCS08_29 ≝
+[
+  quadrupleT ???? (anyOP HCS08 ROL) MODE_DIR1 (Byte 〈x3,x9〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 ROL) MODE_INHA (Byte 〈x4,x9〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 ROL) MODE_INHX (Byte 〈x5,x9〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 ROL) MODE_IX1  (Byte 〈x6,x9〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 ROL) MODE_IX0  (Byte 〈x7,x9〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 ROL) MODE_SP1  (Word 〈〈x9,xE〉:〈x6,x9〉〉) 〈x0,x6〉
+].
+
+definition opcode_table_HCS08_30 ≝
+[
+  quadrupleT ???? (anyOP HCS08 ROR) MODE_DIR1 (Byte 〈x3,x6〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 ROR) MODE_INHA (Byte 〈x4,x6〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 ROR) MODE_INHX (Byte 〈x5,x6〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 ROR) MODE_IX1  (Byte 〈x6,x6〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 ROR) MODE_IX0  (Byte 〈x7,x6〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 ROR) MODE_SP1  (Word 〈〈x9,xE〉:〈x6,x6〉〉) 〈x0,x6〉
+].
+
+definition opcode_table_HCS08_31 ≝
+[
+  quadrupleT ???? (anyOP HCS08 SBC) MODE_IMM1 (Byte 〈xA,x2〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HCS08 SBC) MODE_DIR1 (Byte 〈xB,x2〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 SBC) MODE_DIR2 (Byte 〈xC,x2〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 SBC) MODE_IX2  (Byte 〈xD,x2〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 SBC) MODE_IX1  (Byte 〈xE,x2〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 SBC) MODE_IX0  (Byte 〈xF,x2〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 SBC) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,x2〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 SBC) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,x2〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HCS08_32 ≝
+[
+  quadrupleT ???? (anyOP HCS08 STA) MODE_DIR1 (Byte 〈xB,x7〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 STA) MODE_DIR2 (Byte 〈xC,x7〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 STA) MODE_IX2  (Byte 〈xD,x7〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 STA) MODE_IX1  (Byte 〈xE,x7〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 STA) MODE_IX0  (Byte 〈xF,x7〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HCS08 STA) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,x7〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 STA) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,x7〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HCS08_33 ≝
+[
+  quadrupleT ???? (anyOP HCS08 STX) MODE_DIR1 (Byte 〈xB,xF〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 STX) MODE_DIR2 (Byte 〈xC,xF〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 STX) MODE_IX2  (Byte 〈xD,xF〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 STX) MODE_IX1  (Byte 〈xE,xF〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 STX) MODE_IX0  (Byte 〈xF,xF〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HCS08 STX) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,xF〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 STX) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,xF〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HCS08_34 ≝
+[
+  quadrupleT ???? (anyOP HCS08 SUB) MODE_IMM1 (Byte 〈xA,x0〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP HCS08 SUB) MODE_DIR1 (Byte 〈xB,x0〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 SUB) MODE_DIR2 (Byte 〈xC,x0〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 SUB) MODE_IX2  (Byte 〈xD,x0〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 SUB) MODE_IX1  (Byte 〈xE,x0〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 SUB) MODE_IX0  (Byte 〈xF,x0〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 SUB) MODE_SP2  (Word 〈〈x9,xE〉:〈xD,x0〉〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP HCS08 SUB) MODE_SP1  (Word 〈〈x9,xE〉:〈xE,x0〉〉) 〈x0,x4〉
+].
+
+definition opcode_table_HCS08_35 ≝
+[
+  quadrupleT ???? (anyOP HCS08 TST) MODE_DIR1 (Byte 〈x3,xD〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 TST) MODE_INHA (Byte 〈x4,xD〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 TST) MODE_INHX (Byte 〈x5,xD〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP HCS08 TST) MODE_IX1  (Byte 〈x6,xD〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP HCS08 TST) MODE_IX0  (Byte 〈x7,xD〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP HCS08 TST) MODE_SP1  (Word 〈〈x9,xE〉:〈x6,xD〉〉) 〈x0,x5〉
+].
+
+definition opcode_table_HCS08 ≝
+opcode_table_HCS08_1  @ opcode_table_HCS08_2  @ opcode_table_HCS08_3  @ opcode_table_HCS08_4  @
+opcode_table_HCS08_5  @ opcode_table_HCS08_6  @ opcode_table_HCS08_7  @ opcode_table_HCS08_8  @
+opcode_table_HCS08_9  @ opcode_table_HCS08_10 @ opcode_table_HCS08_11 @ opcode_table_HCS08_12 @
+opcode_table_HCS08_13 @ opcode_table_HCS08_14 @ opcode_table_HCS08_15 @ opcode_table_HCS08_16 @
+opcode_table_HCS08_17 @ opcode_table_HCS08_18 @ opcode_table_HCS08_19 @ opcode_table_HCS08_20 @
+opcode_table_HCS08_21 @ opcode_table_HCS08_22 @ opcode_table_HCS08_23 @ opcode_table_HCS08_24 @
+opcode_table_HCS08_25 @ opcode_table_HCS08_26 @ opcode_table_HCS08_27 @ opcode_table_HCS08_28 @
+opcode_table_HCS08_29 @ opcode_table_HCS08_30 @ opcode_table_HCS08_31 @ opcode_table_HCS08_32 @
+opcode_table_HCS08_33 @ opcode_table_HCS08_34 @ opcode_table_HCS08_35.
+
+(* CORRETTEZZA *)
+
+(* HCS08: opcode non implementati come da manuale (byte) *)
+definition HCS08_not_impl_1 ≝
+ [〈x8,xD〉
+ ;〈x9,xE〉
+ ;〈xA,xC〉
+ ].
+
+lemma ok_byte_table_HCS08 : forall_byte8 (λb.
+ (test_not_impl b HCS08_not_impl_1     ⊕ eqb (get_byte_count HCS08 b 0 opcode_table_HCS08) 1) ⊗
+ (⊖ (test_not_impl b HCS08_not_impl_1) ⊕ eqb (get_byte_count HCS08 b 0 opcode_table_HCS08) 0))
+ = true.
+ reflexivity.
+qed.
+
+(* HCS08: opcode non implementati come da manuale (0x9E+byte) *)
+definition HCS08_not_impl_2 ≝
+ [〈x0,x0〉;〈x0,x1〉;〈x0,x2〉;〈x0,x3〉;〈x0,x4〉;〈x0,x5〉;〈x0,x6〉;〈x0,x7〉
+ ;〈x0,x8〉;〈x0,x9〉;〈x0,xA〉;〈x0,xB〉;〈x0,xC〉;〈x0,xD〉;〈x0,xE〉;〈x0,xF〉
+ ;〈x1,x0〉;〈x1,x1〉;〈x1,x2〉;〈x1,x3〉;〈x1,x4〉;〈x1,x5〉;〈x1,x6〉;〈x1,x7〉
+ ;〈x1,x8〉;〈x1,x9〉;〈x1,xA〉;〈x1,xB〉;〈x1,xC〉;〈x1,xD〉;〈x1,xE〉;〈x1,xF〉
+ ;〈x2,x0〉;〈x2,x1〉;〈x2,x2〉;〈x2,x3〉;〈x2,x4〉;〈x2,x5〉;〈x2,x6〉;〈x2,x7〉
+ ;〈x2,x8〉;〈x2,x9〉;〈x2,xA〉;〈x2,xB〉;〈x2,xC〉;〈x2,xD〉;〈x2,xE〉;〈x2,xF〉
+ ;〈x3,x0〉;〈x3,x1〉;〈x3,x2〉;〈x3,x3〉;〈x3,x4〉;〈x3,x5〉;〈x3,x6〉;〈x3,x7〉
+ ;〈x3,x8〉;〈x3,x9〉;〈x3,xA〉;〈x3,xB〉;〈x3,xC〉;〈x3,xD〉;〈x3,xE〉;〈x3,xF〉
+ ;〈x4,x0〉;〈x4,x1〉;〈x4,x2〉;〈x4,x3〉;〈x4,x4〉;〈x4,x5〉;〈x4,x6〉;〈x4,x7〉
+ ;〈x4,x8〉;〈x4,x9〉;〈x4,xA〉;〈x4,xB〉;〈x4,xC〉;〈x4,xD〉;〈x4,xE〉;〈x4,xF〉
+ ;〈x5,x0〉;〈x5,x1〉;〈x5,x2〉;〈x5,x3〉;〈x5,x4〉;〈x5,x5〉;〈x5,x6〉;〈x5,x7〉
+ ;〈x5,x8〉;〈x5,x9〉;〈x5,xA〉;〈x5,xB〉;〈x5,xC〉;〈x5,xD〉;〈x5,xE〉;〈x5,xF〉
+ ;〈x6,x2〉;〈x6,x5〉;〈x6,xE〉
+ ;〈x7,x0〉;〈x7,x1〉;〈x7,x2〉;〈x7,x3〉;〈x7,x4〉;〈x7,x5〉;〈x7,x6〉;〈x7,x7〉
+ ;〈x7,x8〉;〈x7,x9〉;〈x7,xA〉;〈x7,xB〉;〈x7,xC〉;〈x7,xD〉;〈x7,xE〉;〈x7,xF〉
+ ;〈x8,x0〉;〈x8,x1〉;〈x8,x2〉;〈x8,x3〉;〈x8,x4〉;〈x8,x5〉;〈x8,x6〉;〈x8,x7〉
+ ;〈x8,x8〉;〈x8,x9〉;〈x8,xA〉;〈x8,xB〉;〈x8,xC〉;〈x8,xD〉;〈x8,xE〉;〈x8,xF〉
+ ;〈x9,x0〉;〈x9,x1〉;〈x9,x2〉;〈x9,x3〉;〈x9,x4〉;〈x9,x5〉;〈x9,x6〉;〈x9,x7〉
+ ;〈x9,x8〉;〈x9,x9〉;〈x9,xA〉;〈x9,xB〉;〈x9,xC〉;〈x9,xD〉;〈x9,xE〉;〈x9,xF〉
+ ;〈xA,x0〉;〈xA,x1〉;〈xA,x2〉;〈xA,x3〉;〈xA,x4〉;〈xA,x5〉;〈xA,x6〉;〈xA,x7〉;〈xA,x8〉;〈xA,x9〉;〈xA,xA〉;〈xA,xB〉;〈xA,xC〉;〈xA,xD〉;〈xA,xF〉
+ ;〈xB,x0〉;〈xB,x1〉;〈xB,x2〉;〈xB,x3〉;〈xB,x4〉;〈xB,x5〉;〈xB,x6〉;〈xB,x7〉;〈xB,x8〉;〈xB,x9〉;〈xB,xA〉;〈xB,xB〉;〈xB,xC〉;〈xB,xD〉;〈xB,xF〉
+ ;〈xC,x0〉;〈xC,x1〉;〈xC,x2〉;〈xC,x3〉;〈xC,x4〉;〈xC,x5〉;〈xC,x6〉;〈xC,x7〉;〈xC,x8〉;〈xC,x9〉;〈xC,xA〉;〈xC,xB〉;〈xC,xC〉;〈xC,xD〉;〈xC,xF〉
+ ;〈xD,xC〉;〈xD,xD〉
+ ;〈xE,xC〉;〈xE,xD〉
+ ;〈xF,x0〉;〈xF,x1〉;〈xF,x2〉;〈xF,x4〉;〈xF,x5〉;〈xF,x6〉;〈xF,x7〉;〈xF,x8〉;〈xF,x9〉;〈xF,xA〉;〈xF,xB〉;〈xF,xC〉;〈xF,xD〉
+ ].
+
+lemma ok_byte_table_HCS08_2 : forall_byte8 (λb.
+ (test_not_impl b HCS08_not_impl_2     ⊕ eqb (get_word_count HCS08 b 0 opcode_table_HCS08) 1) ⊗
+ (⊖ (test_not_impl b HCS08_not_impl_2) ⊕ eqb (get_word_count HCS08 b 0 opcode_table_HCS08) 0))
+ = true.
+ reflexivity.
+qed.
+
+(*
+lemma ok_OpIm_table_HCS08 :
+ forall_instr_mode (λi:instr_mode.
+ forall_opcode     (λop:opcode.
+  leb (get_OpIm_count HCS08 (anyOP HCS08 op) i 0 opcode_table_HCS08) 1)) = true.
+ reflexivity.
+qed.
+*)
diff --git a/helm/software/matita/library/freescale/table_RS08.ma b/helm/software/matita/library/freescale/table_RS08.ma
new file mode 100644 (file)
index 0000000..ff2f45f
--- /dev/null
@@ -0,0 +1,427 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/table_RS08/".
+
+(*include "/media/VIRTUOSO/freescale/opcode.ma".*)
+include "freescale/opcode.ma".
+
+(* ***************** *)
+(* TABELLA DELL'RS08 *)
+(* ***************** *)
+
+(* definizione come concatenazione finale di liste per velocizzare il parsing *)
+(* ogni riga e' (any_opcode m) (instr_mode) (opcode esadecimale) (#cicli esecuzione) *)
+(* NB: l'uso di any_opcode m + concatenazione finale tutte liste
+       impedisce di introdurre opcode disomogenei (per mcu) *)
+
+definition opcode_table_RS08_1 ≝
+[ 
+  quadrupleT ???? (anyOP RS08 ADC) MODE_IMM1 (Byte 〈xA,x9〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 ADC) MODE_DIR1 (Byte 〈xB,x9〉) 〈x0,x3〉
+].
+
+definition opcode_table_RS08_2 ≝
+[ 
+  quadrupleT ???? (anyOP RS08 ADD) MODE_IMM1     (Byte 〈xA,xB〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 ADD) MODE_DIR1     (Byte 〈xB,xB〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 ADD) (MODE_TNY x0) (Byte 〈x6,x0〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 ADD) (MODE_TNY x1) (Byte 〈x6,x1〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 ADD) (MODE_TNY x2) (Byte 〈x6,x2〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 ADD) (MODE_TNY x3) (Byte 〈x6,x3〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 ADD) (MODE_TNY x4) (Byte 〈x6,x4〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 ADD) (MODE_TNY x5) (Byte 〈x6,x5〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 ADD) (MODE_TNY x6) (Byte 〈x6,x6〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 ADD) (MODE_TNY x7) (Byte 〈x6,x7〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 ADD) (MODE_TNY x8) (Byte 〈x6,x8〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 ADD) (MODE_TNY x9) (Byte 〈x6,x9〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 ADD) (MODE_TNY xA) (Byte 〈x6,xA〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 ADD) (MODE_TNY xB) (Byte 〈x6,xB〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 ADD) (MODE_TNY xC) (Byte 〈x6,xC〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 ADD) (MODE_TNY xD) (Byte 〈x6,xD〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 ADD) (MODE_TNY xE) (Byte 〈x6,xE〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 ADD) (MODE_TNY xF) (Byte 〈x6,xF〉) 〈x0,x3〉
+].
+
+definition opcode_table_RS08_3 ≝
+[ 
+  quadrupleT ???? (anyOP RS08 AND) MODE_IMM1 (Byte 〈xA,x4〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 AND) MODE_DIR1 (Byte 〈xB,x4〉) 〈x0,x3〉
+].
+
+definition opcode_table_RS08_4 ≝
+[
+  quadrupleT ???? (anyOP RS08 ASL) MODE_INHA (Byte 〈x4,x8〉) 〈x0,x1〉
+].
+
+definition opcode_table_RS08_5 ≝
+[
+  quadrupleT ???? (anyOP RS08 BRA) MODE_IMM1 (Byte 〈x3,x0〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 BCC) MODE_IMM1 (Byte 〈x3,x4〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 BCS) MODE_IMM1 (Byte 〈x3,x5〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 BNE) MODE_IMM1 (Byte 〈x3,x6〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 BEQ) MODE_IMM1 (Byte 〈x3,x7〉) 〈x0,x3〉
+].
+
+definition opcode_table_RS08_6 ≝
+[
+  quadrupleT ???? (anyOP RS08 BSETn) (MODE_DIRn o0) (Byte 〈x1,x0〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BCLRn) (MODE_DIRn o0) (Byte 〈x1,x1〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BSETn) (MODE_DIRn o1) (Byte 〈x1,x2〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BCLRn) (MODE_DIRn o1) (Byte 〈x1,x3〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BSETn) (MODE_DIRn o2) (Byte 〈x1,x4〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BCLRn) (MODE_DIRn o2) (Byte 〈x1,x5〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BSETn) (MODE_DIRn o3) (Byte 〈x1,x6〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BCLRn) (MODE_DIRn o3) (Byte 〈x1,x7〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BSETn) (MODE_DIRn o4) (Byte 〈x1,x8〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BCLRn) (MODE_DIRn o4) (Byte 〈x1,x9〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BSETn) (MODE_DIRn o5) (Byte 〈x1,xA〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BCLRn) (MODE_DIRn o5) (Byte 〈x1,xB〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BSETn) (MODE_DIRn o6) (Byte 〈x1,xC〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BCLRn) (MODE_DIRn o6) (Byte 〈x1,xD〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BSETn) (MODE_DIRn o7) (Byte 〈x1,xE〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BCLRn) (MODE_DIRn o7) (Byte 〈x1,xF〉) 〈x0,x5〉
+].
+
+definition opcode_table_RS08_7 ≝
+[
+  quadrupleT ???? (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x0〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o0) (Byte 〈x0,x1〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x2〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o1) (Byte 〈x0,x3〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x4〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o2) (Byte 〈x0,x5〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x6〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o3) (Byte 〈x0,x7〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x8〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o4) (Byte 〈x0,x9〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xA〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o5) (Byte 〈x0,xB〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xC〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o6) (Byte 〈x0,xD〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BRSETn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xE〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 BRCLRn) (MODE_DIRn_and_IMM1 o7) (Byte 〈x0,xF〉) 〈x0,x5〉
+].
+
+definition opcode_table_RS08_8 ≝
+[
+  quadrupleT ???? (anyOP RS08 CLC ) MODE_INH (Byte 〈x3,x8〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP RS08 SEC ) MODE_INH (Byte 〈x3,x9〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP RS08 SLA ) MODE_INH (Byte 〈x4,x2〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP RS08 SHA ) MODE_INH (Byte 〈x4,x5〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP RS08 NOP ) MODE_INH (Byte 〈xA,xC〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP RS08 STOP) MODE_INH (Byte 〈xA,xE〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 WAIT) MODE_INH (Byte 〈xA,xF〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 RTS ) MODE_INH (Byte 〈xB,xE〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 BGND) MODE_INH (Byte 〈xB,xF〉) 〈x0,x5〉
+].
+
+definition opcode_table_RS08_9 ≝
+[
+  quadrupleT ???? (anyOP RS08 CBEQA) MODE_DIR1_and_IMM1 (Byte 〈x3,x1〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 CBEQA) MODE_IMM1_and_IMM1 (Byte 〈x4,x1〉) 〈x0,x4〉
+].
+
+definition opcode_table_RS08_10 ≝
+[
+  quadrupleT ???? (anyOP RS08 CLR) MODE_DIR1      (Byte 〈x3,xF〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 CLR) MODE_INHA      (Byte 〈x4,xF〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t00) (Byte 〈x8,x0〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t01) (Byte 〈x8,x1〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t02) (Byte 〈x8,x2〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t03) (Byte 〈x8,x3〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t04) (Byte 〈x8,x4〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t05) (Byte 〈x8,x5〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t06) (Byte 〈x8,x6〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t07) (Byte 〈x8,x7〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t08) (Byte 〈x8,x8〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t09) (Byte 〈x8,x9〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t0A) (Byte 〈x8,xA〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t0B) (Byte 〈x8,xB〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t0C) (Byte 〈x8,xC〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t0D) (Byte 〈x8,xD〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t0E) (Byte 〈x8,xE〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t0F) (Byte 〈x8,xF〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t10) (Byte 〈x9,x0〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t11) (Byte 〈x9,x1〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t12) (Byte 〈x9,x2〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t13) (Byte 〈x9,x3〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t14) (Byte 〈x9,x4〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t15) (Byte 〈x9,x5〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t16) (Byte 〈x9,x6〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t17) (Byte 〈x9,x7〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t18) (Byte 〈x9,x8〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t19) (Byte 〈x9,x9〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t1A) (Byte 〈x9,xA〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t1B) (Byte 〈x9,xB〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t1C) (Byte 〈x9,xC〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t1D) (Byte 〈x9,xD〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t1E) (Byte 〈x9,xE〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CLR) (MODE_SRT t1F) (Byte 〈x9,xF〉) 〈x0,x2〉
+].
+
+definition opcode_table_RS08_11 ≝
+[
+  quadrupleT ???? (anyOP RS08 CMP) MODE_IMM1 (Byte 〈xA,x1〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 CMP) MODE_DIR1 (Byte 〈xB,x1〉) 〈x0,x3〉
+].
+
+definition opcode_table_RS08_12 ≝
+[
+  quadrupleT ???? (anyOP RS08 COM) MODE_INHA (Byte 〈x4,x3〉) 〈x0,x1〉
+].
+
+definition opcode_table_RS08_13 ≝
+[
+  quadrupleT ???? (anyOP RS08 DBNZ) MODE_DIR1_and_IMM1 (Byte 〈x3,xB〉) 〈x0,x7〉
+; quadrupleT ???? (anyOP RS08 DBNZ) MODE_INHA_and_IMM1 (Byte 〈x4,xB〉) 〈x0,x4〉
+].
+
+definition opcode_table_RS08_14 ≝
+[
+  quadrupleT ???? (anyOP RS08 DEC) MODE_DIR1     (Byte 〈x3,xA〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 DEC) MODE_INHA     (Byte 〈x4,xA〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP RS08 DEC) (MODE_TNY x0) (Byte 〈x5,x0〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 DEC) (MODE_TNY x1) (Byte 〈x5,x1〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 DEC) (MODE_TNY x2) (Byte 〈x5,x2〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 DEC) (MODE_TNY x3) (Byte 〈x5,x3〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 DEC) (MODE_TNY x4) (Byte 〈x5,x4〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 DEC) (MODE_TNY x5) (Byte 〈x5,x5〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 DEC) (MODE_TNY x6) (Byte 〈x5,x6〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 DEC) (MODE_TNY x7) (Byte 〈x5,x7〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 DEC) (MODE_TNY x8) (Byte 〈x5,x8〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 DEC) (MODE_TNY x9) (Byte 〈x5,x9〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 DEC) (MODE_TNY xA) (Byte 〈x5,xA〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 DEC) (MODE_TNY xB) (Byte 〈x5,xB〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 DEC) (MODE_TNY xC) (Byte 〈x5,xC〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 DEC) (MODE_TNY xD) (Byte 〈x5,xD〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 DEC) (MODE_TNY xE) (Byte 〈x5,xE〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 DEC) (MODE_TNY xF) (Byte 〈x5,xF〉) 〈x0,x4〉
+].
+
+definition opcode_table_RS08_15 ≝
+[ 
+  quadrupleT ???? (anyOP RS08 EOR) MODE_IMM1 (Byte 〈xA,x8〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 EOR) MODE_DIR1 (Byte 〈xB,x8〉) 〈x0,x3〉
+].
+
+definition opcode_table_RS08_16 ≝
+[
+  quadrupleT ???? (anyOP RS08 INC) MODE_DIR1     (Byte 〈x3,xC〉) 〈x0,x5〉
+; quadrupleT ???? (anyOP RS08 INC) MODE_INHA     (Byte 〈x4,xC〉) 〈x0,x1〉
+; quadrupleT ???? (anyOP RS08 INC) (MODE_TNY x0) (Byte 〈x2,x0〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 INC) (MODE_TNY x1) (Byte 〈x2,x1〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 INC) (MODE_TNY x2) (Byte 〈x2,x2〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 INC) (MODE_TNY x3) (Byte 〈x2,x3〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 INC) (MODE_TNY x4) (Byte 〈x2,x4〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 INC) (MODE_TNY x5) (Byte 〈x2,x5〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 INC) (MODE_TNY x6) (Byte 〈x2,x6〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 INC) (MODE_TNY x7) (Byte 〈x2,x7〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 INC) (MODE_TNY x8) (Byte 〈x2,x8〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 INC) (MODE_TNY x9) (Byte 〈x2,x9〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 INC) (MODE_TNY xA) (Byte 〈x2,xA〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 INC) (MODE_TNY xB) (Byte 〈x2,xB〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 INC) (MODE_TNY xC) (Byte 〈x2,xC〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 INC) (MODE_TNY xD) (Byte 〈x2,xD〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 INC) (MODE_TNY xE) (Byte 〈x2,xE〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 INC) (MODE_TNY xF) (Byte 〈x2,xF〉) 〈x0,x4〉
+].
+
+definition opcode_table_RS08_17 ≝
+[
+  quadrupleT ???? (anyOP RS08 JMP) MODE_DIR2 (Byte 〈xB,xC〉) 〈x0,x4〉
+].
+
+definition opcode_table_RS08_18 ≝
+[
+  quadrupleT ???? (anyOP RS08 BSR) MODE_IMM1 (Byte 〈xA,xD〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 JSR) MODE_DIR2 (Byte 〈xB,xD〉) 〈x0,x4〉
+].
+
+definition opcode_table_RS08_19 ≝
+[
+  quadrupleT ???? (anyOP RS08 LDA) MODE_IMM1      (Byte 〈xA,x6〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 LDA) MODE_DIR1      (Byte 〈xB,x6〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t00) (Byte 〈xC,x0〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t01) (Byte 〈xC,x1〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t02) (Byte 〈xC,x2〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t03) (Byte 〈xC,x3〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t04) (Byte 〈xC,x4〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t05) (Byte 〈xC,x5〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t06) (Byte 〈xC,x6〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t07) (Byte 〈xC,x7〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t08) (Byte 〈xC,x8〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t09) (Byte 〈xC,x9〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t0A) (Byte 〈xC,xA〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t0B) (Byte 〈xC,xB〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t0C) (Byte 〈xC,xC〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t0D) (Byte 〈xC,xD〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t0E) (Byte 〈xC,xE〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t0F) (Byte 〈xC,xF〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t10) (Byte 〈xD,x0〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t11) (Byte 〈xD,x1〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t12) (Byte 〈xD,x2〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t13) (Byte 〈xD,x3〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t14) (Byte 〈xD,x4〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t15) (Byte 〈xD,x5〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t16) (Byte 〈xD,x6〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t17) (Byte 〈xD,x7〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t18) (Byte 〈xD,x8〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t19) (Byte 〈xD,x9〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t1A) (Byte 〈xD,xA〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t1B) (Byte 〈xD,xB〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t1C) (Byte 〈xD,xC〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t1D) (Byte 〈xD,xD〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t1E) (Byte 〈xD,xE〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 LDA) (MODE_SRT t1F) (Byte 〈xD,xF〉) 〈x0,x3〉
+].
+
+definition opcode_table_RS08_20 ≝
+[
+  quadrupleT ???? (anyOP RS08 LSR) MODE_INHA (Byte 〈x4,x4〉) 〈x0,x1〉
+].
+
+definition opcode_table_RS08_21 ≝
+[
+  quadrupleT ???? (anyOP RS08 MOV) MODE_IMM1_to_DIR1 (Byte 〈x3,xE〉) 〈x0,x4〉
+; quadrupleT ???? (anyOP RS08 MOV) MODE_DIR1_to_DIR1 (Byte 〈x4,xE〉) 〈x0,x5〉
+].
+
+definition opcode_table_RS08_22 ≝
+[ 
+  quadrupleT ???? (anyOP RS08 ORA) MODE_IMM1 (Byte 〈xA,xA〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 ORA) MODE_DIR1 (Byte 〈xB,xA〉) 〈x0,x3〉
+].
+
+definition opcode_table_RS08_23 ≝
+[
+  quadrupleT ???? (anyOP RS08 ROL) MODE_INHA (Byte 〈x4,x9〉) 〈x0,x1〉
+].
+
+definition opcode_table_RS08_24 ≝
+[
+  quadrupleT ???? (anyOP RS08 ROR) MODE_INHA (Byte 〈x4,x6〉) 〈x0,x1〉
+].
+
+definition opcode_table_RS08_25 ≝
+[
+  quadrupleT ???? (anyOP RS08 SBC) MODE_IMM1 (Byte 〈xA,x2〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 SBC) MODE_DIR1 (Byte 〈xB,x2〉) 〈x0,x3〉
+].
+
+definition opcode_table_RS08_26 ≝
+[
+  quadrupleT ???? (anyOP RS08 STA) MODE_DIR1      (Byte 〈xB,x7〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t00) (Byte 〈xE,x0〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t01) (Byte 〈xE,x1〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t02) (Byte 〈xE,x2〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t03) (Byte 〈xE,x3〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t04) (Byte 〈xE,x4〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t05) (Byte 〈xE,x5〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t06) (Byte 〈xE,x6〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t07) (Byte 〈xE,x7〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t08) (Byte 〈xE,x8〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t09) (Byte 〈xE,x9〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t0A) (Byte 〈xE,xA〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t0B) (Byte 〈xE,xB〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t0C) (Byte 〈xE,xC〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t0D) (Byte 〈xE,xD〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t0E) (Byte 〈xE,xE〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t0F) (Byte 〈xE,xF〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t10) (Byte 〈xF,x0〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t11) (Byte 〈xF,x1〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t12) (Byte 〈xF,x2〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t13) (Byte 〈xF,x3〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t14) (Byte 〈xF,x4〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t15) (Byte 〈xF,x5〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t16) (Byte 〈xF,x6〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t17) (Byte 〈xF,x7〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t18) (Byte 〈xF,x8〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t19) (Byte 〈xF,x9〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t1A) (Byte 〈xF,xA〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t1B) (Byte 〈xF,xB〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t1C) (Byte 〈xF,xC〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t1D) (Byte 〈xF,xD〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t1E) (Byte 〈xF,xE〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 STA) (MODE_SRT t1F) (Byte 〈xF,xF〉) 〈x0,x2〉
+].
+
+definition opcode_table_RS08_27 ≝
+[ 
+  quadrupleT ???? (anyOP RS08 SUB) MODE_IMM1     (Byte 〈xA,x0〉) 〈x0,x2〉
+; quadrupleT ???? (anyOP RS08 SUB) MODE_DIR1     (Byte 〈xB,x0〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 SUB) (MODE_TNY x0) (Byte 〈x7,x0〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 SUB) (MODE_TNY x1) (Byte 〈x7,x1〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 SUB) (MODE_TNY x2) (Byte 〈x7,x2〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 SUB) (MODE_TNY x3) (Byte 〈x7,x3〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 SUB) (MODE_TNY x4) (Byte 〈x7,x4〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 SUB) (MODE_TNY x5) (Byte 〈x7,x5〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 SUB) (MODE_TNY x6) (Byte 〈x7,x6〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 SUB) (MODE_TNY x7) (Byte 〈x7,x7〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 SUB) (MODE_TNY x8) (Byte 〈x7,x8〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 SUB) (MODE_TNY x9) (Byte 〈x7,x9〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 SUB) (MODE_TNY xA) (Byte 〈x7,xA〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 SUB) (MODE_TNY xB) (Byte 〈x7,xB〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 SUB) (MODE_TNY xC) (Byte 〈x7,xC〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 SUB) (MODE_TNY xD) (Byte 〈x7,xD〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 SUB) (MODE_TNY xE) (Byte 〈x7,xE〉) 〈x0,x3〉
+; quadrupleT ???? (anyOP RS08 SUB) (MODE_TNY xF) (Byte 〈x7,xF〉) 〈x0,x3〉
+].
+
+definition opcode_table_RS08 ≝
+opcode_table_RS08_1  @ opcode_table_RS08_2  @ opcode_table_RS08_3  @ opcode_table_RS08_4  @
+opcode_table_RS08_5  @ opcode_table_RS08_6  @ opcode_table_RS08_7  @ opcode_table_RS08_8  @
+opcode_table_RS08_9  @ opcode_table_RS08_10 @ opcode_table_RS08_11 @ opcode_table_RS08_12 @
+opcode_table_RS08_13 @ opcode_table_RS08_14 @ opcode_table_RS08_15 @ opcode_table_RS08_16 @
+opcode_table_RS08_17 @ opcode_table_RS08_18 @ opcode_table_RS08_19 @ opcode_table_RS08_20 @
+opcode_table_RS08_21 @ opcode_table_RS08_22 @ opcode_table_RS08_23 @ opcode_table_RS08_24 @
+opcode_table_RS08_25 @ opcode_table_RS08_26 @ opcode_table_RS08_27.
+
+(* CORRETTEZZA *)
+
+(* RS08: opcode non implementati come da manuale *)
+definition RS08_not_impl ≝
+ [〈x3,x2〉;〈x3,x3〉;〈x3,xD〉
+ ;〈x4,x0〉;〈x4,x7〉;〈x4,xD〉
+ ;〈xA,x3〉;〈xA,x5〉;〈xA,x7〉
+ ;〈xB,x3〉;〈xB,x5〉
+ ].
+
+lemma ok_byte_table_RS08 : forall_byte8 (λb.
+ (test_not_impl b RS08_not_impl     ⊕ eqb (get_byte_count RS08 b 0 opcode_table_RS08) 1) ⊗
+ (⊖ (test_not_impl b RS08_not_impl) ⊕ eqb (get_byte_count RS08 b 0 opcode_table_RS08) 0))
+ = true.
+ reflexivity.
+qed.
+
+(*
+lemma ok_OpIm_table_RS08 :
+ forall_instr_mode (λi:instr_mode.
+ forall_opcode     (λop:opcode.
+  leb (get_OpIm_count RS08 (anyOP RS08 op) i 0 opcode_table_RS08) 1)) = true.
+ reflexivity.
+qed.
+*)
diff --git a/helm/software/matita/library/freescale/tests.old b/helm/software/matita/library/freescale/tests.old
new file mode 100644 (file)
index 0000000..17766c7
--- /dev/null
@@ -0,0 +1,311 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/tests/".
+
+(*include "/media/VIRTUOSO/freescale/micro_tests.ma".*)
+include "freescale/micro_tests.ma".
+
+(* 
+   RAM indirizzabile in modalita' diretta da usare per X,Y,Z
+   A=X*Y (parte low) con [0x0020-0x004F] X ≝ [0x0020] Y ≝ [0x0021] Z ≝ [0x0022]
+*)
+definition test_mult_source_RS08 : list byte8 ≝
+let m ≝ RS08 in source_to_byte8 m (
+(* [0x3800] Z <- 0      3clk *) (compile m ? CLR (maDIR1 〈x2,x2〉) I) @
+(* [0x3802] (l1) A <- Y 3clk *) (compile m ? LDA (maDIR1 〈x2,x1〉) I) @
+(* [0x3804] A=0 goto l2 3clk *) (compile m ? BEQ (maIMM1 〈x0,xA〉) I) @
+(* [0x3806] A <- Z      3clk *) (compile m ? LDA (maDIR1 〈x2,x2〉) I) @
+(* [0x3808] Y --        5clk *) (compile m ? DEC (maDIR1 〈x2,x1〉) I) @
+(* [0x380A] A += X      3clk *) (compile m ? ADD (maDIR1 〈x2,x0〉) I) @
+(* [0x380C] Z <- A      3clk *) (compile m ? STA (maDIR1 〈x2,x2〉) I) @
+(* [0x380E] goto l1     3clk *) (compile m ? BRA (maIMM1 〈xF,x2〉) I) @
+(* [0x3810] (l2) A <- Z 3clk *) (compile m ? LDA (maDIR1 〈x2,x2〉) I)
+ ).
+
+(* ************** *)
+(* *****TODO***** *)
+(* ************** *)
+
+definition mult_memory ≝
+ λx,y.λa:addr.
+     match leb a 29 with
+      [ true ⇒ nth ? mult_source 〈x0, x0〉 a
+      | false ⇒
+         match eqb a 30 with
+          [ true ⇒ x
+          | false ⇒ y
+          ]
+      ].
+
+definition mult_status ≝
+ λx,y.
+  mk_status 〈x0, x0〉 0 0 false false (mult_memory x y) 0.
+
+notation " 'M' \sub (x y)" non associative with precedence 80 for 
+ @{ 'memory $x $y }.
+interpretation "mult_memory" 'memory x y = 
+  (cic:/matita/freescale/test/mult_memory.con x y).
+
+notation " 'M' \sub (x y) \nbsp a" non associative with precedence 80 for 
+ @{ 'memory4 $x $y $a }.
+interpretation "mult_memory4" 'memory4 x y a = 
+  (cic:/matita/freescale/test/mult_memory.con x y a).
+  
+notation " \Sigma \sub (x y)" non associative with precedence 80 for 
+ @{ 'status $x $y }.
+
+interpretation "mult_status" 'status x y =
+  (cic:/matita/freescale/test/mult_status.con x y).
+
+lemma test_O_O:
+  let i ≝ 14 in
+  let s ≝ execute (mult_status 〈x0, x0〉 〈x0, x0〉) i in
+   pc s = 20 ∧ mem s 32 = byte_of_nat 0.
+ split;
+ reflexivity.
+qed.
+
+lemma test_0_2:
+  let x ≝ 〈x0, x0〉 in
+  let y ≝ 〈x0, x2〉 in
+  let i ≝ 14 + 23 * nat_of_byte y in
+  let s ≝ execute (mult_status x y) i in
+   pc s = 20 ∧ mem s 32 = plusbytenc x x.
+ intros;
+ split;
+ reflexivity.
+qed.
+
+lemma test_x_1:
+ ∀x.
+  let y ≝ 〈x0, x1〉 in
+  let i ≝ 14 + 23 * nat_of_byte y in
+  let s ≝ execute (mult_status x y) i in
+   pc s = 20 ∧ mem s 32 = x.
+ intros;
+ split;
+  [ reflexivity
+  | change in ⊢ (? ? % ?) with (plusbytenc 〈x0, x0〉 x);
+    rewrite > plusbytenc_O_x;
+    reflexivity
+  ].
+qed.
+
+lemma test_x_2:
+ ∀x.
+  let y ≝ 〈x0, x2〉 in
+  let i ≝ 14 + 23 * nat_of_byte y in
+  let s ≝ execute (mult_status x y) i in
+   pc s = 20 ∧ mem s 32 = plusbytenc x x.
+ intros;
+ split;
+  [ reflexivity
+  | change in ⊢ (? ? % ?) with
+     (plusbytenc (plusbytenc 〈x0, x0〉 x) x);
+    rewrite > plusbytenc_O_x;
+    reflexivity
+  ].
+qed.
+
+lemma loop_invariant':
+ ∀x,y:byte.∀j:nat. j ≤ y →
+  execute (mult_status x y) (5 + 23*j)
+   =
+    mk_status (byte_of_nat (x * j)) 4 0 (eqbyte 〈x0, x0〉 (byte_of_nat (x*j)))
+     (plusbytec (byte_of_nat (x*pred j)) x)
+     (update (update (update (mult_memory x y) 30 x) 31 (byte_of_nat (y - j))) 32
+      (byte_of_nat (x * j)))
+     0.
+ intros 3;
+ elim j;
+  [ do 2 (rewrite < times_n_O);
+    apply status_eq;
+    [1,2,3,4,7: normalize; reflexivity
+    | rewrite > eq_plusbytec_x0_x0_x_false;
+      normalize;
+      reflexivity 
+    | intro;
+      rewrite < minus_n_O;
+      normalize in ⊢ (? ? (? (? ? %) ?) ?);
+      change in ⊢ (? ? % ?) with (update (mult_memory x y) 32 〈x0, x0〉 a);
+      simplify in ⊢ (? ? ? %);
+      change in ⊢ (? ? ? (? (? (? ? ? %) ? ?) ? ? ?)) with (mult_memory x y 30);
+      rewrite > byte_of_nat_nat_of_byte;
+      change in ⊢ (? ? ? (? (? ? ? %) ? ? ?)) with (mult_memory x y 31);
+      apply inj_update;
+      intro;
+      rewrite > (eq_update_s_a_sa (update (mult_memory x y) 30 (mult_memory x y 30))
+       31 a);
+      rewrite > eq_update_s_a_sa;
+      reflexivity
+    ]
+  | cut (5 + 23 * S n = 5 + 23 * n + 23);
+    [ rewrite > Hcut; clear Hcut;
+      rewrite > breakpoint;
+      rewrite > H; clear H;
+      [2: apply le_S_S_to_le;
+        apply le_S;
+        apply H1
+      | cut (∃z.y-n=S z ∧ z < 255);
+         [ elim Hcut; clear Hcut;
+           elim H; clear H;
+           rewrite > H2;
+           (* instruction LDAd *)
+           change in ⊢ (? ? (? ? %) ?) with (3+20);
+           rewrite > breakpoint in ⊢ (? ? % ?);
+           whd in ⊢ (? ? (? % ?) ?);
+           normalize in ⊢ (? ? (? (? ? % ? ? ? ? ?) ?) ?);
+           change in ⊢ (? ? (? (? % ? ? ? ? ? ?) ?) ?)
+            with (byte_of_nat (S a));
+           change in ⊢ (? ? (? (? ? ? ? (? ? %) ? ? ?) ?) ?) with
+            (byte_of_nat (S a));
+           (* instruction BEQ *)
+           change in ⊢ (? ? (? ? %) ?) with (3+17);
+           rewrite > breakpoint in ⊢ (? ? % ?);
+           whd in ⊢ (? ? (? % ?) ?);
+           letin K ≝ (eq_eqbyte_x0_x0_byte_of_nat_S_false ? H3); clearbody K;
+           rewrite > K; clear K;
+           simplify in ⊢ (? ? (? (? ? % ? ? ? ? ?) ?) ?);
+           (* instruction LDAd *)
+           change in ⊢ (? ? (? ? %) ?) with (3+14);
+           rewrite > breakpoint in ⊢ (? ? % ?);
+           whd in ⊢ (? ? (? % ?) ?);
+           change in ⊢ (? ? (? (? % ? ? ? ? ? ?) ?) ?) with (byte_of_nat (x*n));
+           normalize in ⊢ (? ? (? (? ? % ? ? ? ? ?) ?) ?);
+           change in ⊢ (? ? (? (? ? ? ? % ? ? ?) ?) ?) with (eqbyte 〈x0, x0〉 (byte_of_nat (x*n)));
+           (* instruction DECd *)
+           change in ⊢ (? ? (? ? %) ?) with (5+9);
+           rewrite > breakpoint in ⊢ (? ? % ?);
+           whd in ⊢ (? ? (? % ?) ?);
+           change in ⊢ (? ? (? (? ? ? ? (? ? %) ? ? ?) ?) ?) with (bpred (byte_of_nat (S a)));
+           rewrite > (eq_bpred_S_a_a ? H3);
+           normalize in ⊢ (? ? (? (? ? % ? ? ? ? ?) ?) ?);
+           normalize in ⊢ (? ? (? (? ? ? ? ? ? (? ? % ?) ?) ?) ?);
+           cut (y - S n = a);
+            [2: rewrite > eq_minus_S_pred;
+                rewrite > H2;
+                reflexivity | ];
+           rewrite < Hcut; clear Hcut; clear H3; clear H2; clear a;          
+           (* instruction ADDd *)
+           change in ⊢ (? ? (? ? %) ?) with (3+6);
+           rewrite > breakpoint in ⊢ (? ? % ?);
+           whd in ⊢ (? ? (? % ?) ?);
+           change in ⊢ (? ? (? (? % ? ? ? ? ? ?) ?) ?) with
+            (plusbytenc (byte_of_nat (x*n)) x);
+           change in ⊢ (? ? (? (? ? ? ? (? ? %) ? ? ?) ?) ?) with
+            (plusbytenc (byte_of_nat (x*n)) x);
+           normalize in ⊢ (? ? (? (? ? % ? ? ? ? ?) ?) ?);
+           change in ⊢ (? ? (? (? ? ? ? ? % ? ?) ?) ?)
+            with (plusbytec (byte_of_nat (x*n)) x);
+           rewrite > plusbytenc_S;
+           (* instruction STAd *)
+           rewrite > (breakpoint ? 3 3);
+           whd in ⊢ (? ? (? % ?) ?);
+           normalize in ⊢ (? ? (? (? ? % ? ? ? ? ?) ?) ?);
+           (* instruction BRA *)
+           whd in ⊢ (? ? % ?);
+           normalize in ⊢ (? ? (? ? % ? ? ? ? ?) ?);
+           rewrite < pred_Sn;        
+           apply status_eq;
+            [1,2,3,4,7: normalize; reflexivity
+            | change with (plusbytec #(x*n) x = plusbytec #(x*n) x);
+              reflexivity
+            |6: intro;
+              simplify in ⊢ (? ? ? %);
+              normalize in ⊢ (? ? (? (? ? ? ? ? ? (? ? (? %) ?) ?) ?) ?);
+              change in ⊢ (? ? % ?) with
+               ((mult_memory x y){30↦x}{31↦#(S (y-S n))}{32↦#(x*n)}{31↦#(y-S n)}
+                {〈x2,x0〉↦ #(x*S n)} a);
+              apply inj_update;
+              intro;
+              apply inj_update;
+              intro;
+              rewrite > not_eq_a_b_to_eq_update_a_b; [2: apply H | ];
+              rewrite > not_eq_a_b_to_eq_update_a_b;
+               [ reflexivity
+               | assumption
+               ]
+            ]
+         | exists;
+            [ apply (y - S n)
+            | split;
+               [ rewrite < (minus_S_S y n);
+                 apply (minus_Sn_m (nat_of_byte y) (S n) H1)
+               | letin K ≝ (lt_nat_of_byte_256 y); clearbody K;
+                 letin K' ≝ (lt_minus_m y (S n) ? ?); clearbody K';
+                 [ apply (lt_to_le_to_lt O (S n) (nat_of_byte y) ? ?);
+                   autobatch
+                 | autobatch
+                 | autobatch
+                 ]
+               ]
+            ]
+         ]
+      ]
+    | rewrite > associative_plus;
+      rewrite < times_n_Sm;
+      rewrite > sym_plus in ⊢ (? ? ? (? ? %));
+      reflexivity
+    ] 
+  ]   
+qed.
+
+
+theorem test_x_y:
+ ∀x,y:byte.
+  let i ≝ 14 + 23 * y in
+   execute (mult_status x y) i =
+    mk_status (#(x*y)) 20 0
+     (eqbyte 〈x0, x0〉 (#(x*y)))
+     (plusbytec (byte_of_nat (x*pred y)) x)
+     (update
+       (update (mult_memory x y) 31 〈x0, x0〉)
+       32 (byte_of_nat (x*y)))
+     0.
+ intros;
+ cut (14 + 23 * y = 5 + 23*y + 9);
+  [2: autobatch paramodulation;
+  | rewrite > Hcut; (* clear Hcut; *)
+    rewrite > (breakpoint (mult_status x y) (5 + 23*y) 9);
+    rewrite > loop_invariant';
+     [2: apply le_n
+     | rewrite < minus_n_n;
+       apply status_eq;
+        [1,2,3,4,5,7: normalize; reflexivity
+        | intro;
+          simplify in ⊢ (? ? ? %);
+          change in ⊢ (? ? % ?) with
+           (update (update (update (mult_memory x y) 30 x) 31 (byte_of_nat O)) 32
+(byte_of_nat (nat_of_byte x*nat_of_byte y)) a);
+          repeat (apply inj_update; intro);
+          apply (eq_update_s_a_sa ? 30)
+        ]
+     ]
+  ].
+qed.
diff --git a/helm/software/matita/library/freescale/translation.ma b/helm/software/matita/library/freescale/translation.ma
new file mode 100644 (file)
index 0000000..dae2b5f
--- /dev/null
@@ -0,0 +1,241 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/translation/".
+
+(*include "/media/VIRTUOSO/freescale/table_HC05.ma".*)
+include "freescale/table_HC05.ma".
+(*include "/media/VIRTUOSO/freescale/table_HC08.ma".*)
+include "freescale/table_HC08.ma".
+(*include "/media/VIRTUOSO/freescale/table_HCS08.ma".*)
+include "freescale/table_HCS08.ma".
+(*include "/media/VIRTUOSO/freescale/table_RS08.ma".*)
+include "freescale/table_RS08.ma".
+
+(* ***************************** *)
+(* TRADUZIONE ESADECIMALE → INFO *)
+(* ***************************** *)
+
+(* accesso alla tabella della ALU prescelta *)
+definition opcode_table ≝
+λm:mcu_type.
+ match m
+  return λm:mcu_type.list (Prod4T (any_opcode m) instr_mode byte8_or_word16 byte8)
+ with
+  [ HC05  ⇒ opcode_table_HC05
+  | HC08  ⇒ opcode_table_HC08
+  | HCS08 ⇒ opcode_table_HCS08
+  | RS08  ⇒ opcode_table_RS08
+ ].
+
+(* traduzione mcu+esadecimale → info *)
+(* NB: la ricerca per byte non matcha con una word con lo stesso byte superiore uguale *)
+(* NB: per matchare una word (precode+code) bisogna passare una word *)
+(* NB: il risultato e' sempre un'opzione, NON esiste un dummy opcode tipo UNKNOWN/ILLEGAL *)
+definition full_info_of_word16 ≝
+λm:mcu_type.λborw:byte8_or_word16.
+ let rec aux param ≝ match param with
+  [ nil ⇒ None ?
+  | cons hd tl ⇒
+   match thd4T ???? hd with
+    [ Byte b ⇒ match borw with
+     [ Byte borw' ⇒ match eq_b8 borw' b with
+      [ true ⇒ Some ? hd | false ⇒ aux tl ]
+     | Word _ ⇒ aux tl ]
+    | Word w ⇒ match borw with
+     [ Byte _ ⇒ aux tl
+     | Word borw' ⇒ match eq_w16 borw' w with
+      [ true ⇒ Some ? hd | false ⇒ aux tl ]            
+    ]]] in aux (opcode_table m).
+
+(* ******************************************************* *)
+(* TRADUZIONE MCU+OPCODE+MODALITA'+ARGOMENTI → ESADECIMALE *)
+(* ******************************************************* *)
+
+(* introduzione di un tipo byte8 dipendente dall'mcu_type (phantom type) *)
+inductive t_byte8 (m:mcu_type) : Type ≝
+ TByte : byte8 → t_byte8 m.
+
+coercion cic:/matita/freescale/translation/t_byte8.ind#xpointer(1/1/1).
+
+(* introduzione di un tipo dipendente (dalla modalita') per gli argomenti *)
+inductive MA_check : instr_mode → Type ≝
+  maINH              : MA_check MODE_INH
+| maINHA             : MA_check MODE_INHA
+| maINHX             : MA_check MODE_INHX
+| maINHH             : MA_check MODE_INHH
+| maIMM1             : byte8  → MA_check MODE_IMM1
+| maIMM2             : word16 → MA_check MODE_IMM2
+| maDIR1             : byte8  → MA_check MODE_DIR1
+| maDIR2             : word16 → MA_check MODE_DIR2
+| maIX0              : MA_check MODE_IX0
+| maIX1              : byte8  → MA_check MODE_IX1
+| maIX2              : word16 → MA_check MODE_IX2
+| maSP1              : byte8  → MA_check MODE_SP1
+| maSP2              : word16 → MA_check MODE_SP2
+| maDIR1_to_DIR1     : byte8 → byte8 → MA_check MODE_DIR1_to_DIR1
+| maIMM1_to_DIR1     : byte8 → byte8 → MA_check MODE_IMM1_to_DIR1
+| maIX0p_to_DIR1     : byte8 → MA_check MODE_IX0p_to_DIR1
+| maDIR1_to_IX0p     : byte8 → MA_check MODE_DIR1_to_IX0p
+| maINHA_and_IMM1    : byte8 → MA_check MODE_INHA_and_IMM1
+| maINHX_and_IMM1    : byte8 → MA_check MODE_INHX_and_IMM1
+| maIMM1_and_IMM1    : byte8 → byte8 → MA_check MODE_IMM1_and_IMM1
+| maDIR1_and_IMM1    : byte8 → byte8 → MA_check MODE_DIR1_and_IMM1
+| maIX0_and_IMM1     : byte8 → MA_check MODE_IX0_and_IMM1
+| maIX0p_and_IMM1    : byte8 → MA_check MODE_IX0p_and_IMM1
+| maIX1_and_IMM1     : byte8 → byte8 → MA_check MODE_IX1_and_IMM1
+| maIX1p_and_IMM1    : byte8 → byte8 → MA_check MODE_IX1p_and_IMM1
+| maSP1_and_IMM1     : byte8 → byte8 → MA_check MODE_SP1_and_IMM1
+| maDIRn             : ∀o.byte8 → MA_check (MODE_DIRn o)
+| maDIRn_and_IMM1    : ∀o.byte8 → byte8 → MA_check (MODE_DIRn_and_IMM1 o)
+| maTNY              : ∀e.MA_check (MODE_TNY e)
+| maSRT              : ∀t.MA_check (MODE_SRT t)
+.
+
+(* tipo istruzione per unificare in una lista omogenea il sorgente *)
+inductive instruction : Type ≝ 
+instr: ∀i:instr_mode.opcode → MA_check i → instruction.
+
+coercion cic:/matita/freescale/translation/instruction.ind#xpointer(1/1/1).
+
+(* picker: trasforma l'argomento necessario in input a bytes_of_pseudo_instr_mode_param:
+   MA_check i → list (t_byte8 m) *)
+definition args_picker ≝
+λm:mcu_type.λi:instr_mode.λargs:MA_check i.
+ match args with
+  (* inherent: legale se nessun operando *) 
+  [ maINH    ⇒ nil ? 
+  | maINHA   ⇒ nil ? 
+  | maINHX   ⇒ nil ? 
+  | maINHH   ⇒ nil ?
+  (* _0/1/2: legale se nessun operando/1 byte/1 word *)
+  | maIMM1 b ⇒ [ TByte m b ]
+  | maIMM2 w ⇒ [ TByte m (w16h w); TByte m (w16l w) ]
+  | maDIR1 b ⇒ [ TByte m b ]
+  | maDIR2 w ⇒ [ TByte m (w16h w); TByte m (w16l w) ]
+  | maIX0    ⇒ nil ?
+  | maIX1 b  ⇒ [ TByte m b ]
+  | maIX2 w  ⇒ [ TByte m (w16h w); TByte m (w16l w) ]
+  | maSP1 b  ⇒ [ TByte m b ]
+  | maSP2 w  ⇒ [ TByte m (w16h w); TByte m (w16l w) ]
+  (* movimento: legale se 2 operandi byte *)
+  | maDIR1_to_DIR1 b1 b2  ⇒ [ TByte m b1 ; TByte m b2 ]
+  | maIMM1_to_DIR1 b1 b2  ⇒ [ TByte m b1 ; TByte m b2 ]
+  | maIX0p_to_DIR1 b      ⇒ [ TByte m b]
+  | maDIR1_to_IX0p b      ⇒ [ TByte m b]
+  (* cbeq/dbnz: legale se 1/2 operandi byte *)
+  | maINHA_and_IMM1 b     ⇒ [ TByte m b]
+  | maINHX_and_IMM1 b     ⇒ [ TByte m b]
+  | maIMM1_and_IMM1 b1 b2 ⇒ [ TByte m b1 ; TByte m b2 ]
+  | maDIR1_and_IMM1 b1 b2 ⇒ [ TByte m b1 ; TByte m b2 ]
+  | maIX0_and_IMM1  b     ⇒ [ TByte m b]
+  | maIX0p_and_IMM1 b     ⇒ [ TByte m b]
+  | maIX1_and_IMM1  b1 b2 ⇒ [ TByte m b1 ; TByte m b2 ]
+  | maIX1p_and_IMM1 b1 b2 ⇒ [ TByte m b1 ; TByte m b2 ]
+  | maSP1_and_IMM1  b1 b2 ⇒ [ TByte m b1 ; TByte m b2 ]
+  (* DIRn: legale se 1 operando byte *)
+  | maDIRn _ b ⇒ [ TByte m b]
+  (* DIRn_and_IMM1: legale se 2 operandi byte *)
+  | maDIRn_and_IMM1 _ b1 b2 ⇒ [ TByte m b1 ; TByte m b2 ]
+  (* TNY: legale se nessun operando *)
+  | maTNY _ ⇒ nil ?
+  (* SRT: legale se nessun operando *)
+  | maSRT _ ⇒ nil ?
+  ].
+
+(* trasformatore finale: mcu+opcode+instr_mode+args → list (t_byte8 m) *)
+definition bytes_of_pseudo_instr_mode_param ≝
+λm:mcu_type.λo:any_opcode m.λi:instr_mode.λp:MA_check i.
+let rec aux param  ≝ match param with
+ [ nil ⇒ None ? | cons hd tl ⇒
+  match (eqop m o (fst4T ???? hd)) ⊗ (eqim i (snd4T ???? hd)) with
+   [ true ⇒ match thd4T ???? hd with 
+    [ Byte isab ⇒ 
+     Some ? ([ (TByte m isab) ]@(args_picker m i p))
+    | Word isaw ⇒
+     Some ? ([ (TByte m (w16h isaw)) ; (TByte m (w16l isaw)) ]@(args_picker m i p))
+    ]
+   | false ⇒ aux tl ]] in aux (opcode_table m).
+
+(* tipatore degli opcode generici *)
+definition opcode_to_any ≝ λm:mcu_type.λo:opcode.anyOP m o.
+
+(* ****************************** *)
+(* APPROCCIO COMPILAZIONE AL VOLO *)
+(* ****************************** *)
+
+(* ausiliario di compile *)
+definition defined ≝
+ λT:Type.λo:option T.
+  match o with
+   [ None ⇒ False
+   | Some _ ⇒ True
+   ].
+
+(* compila solo se l'intera istruzione+modalita'+argomenti ha riscontro nelle tabelle *)
+definition compile ≝
+λmcu:mcu_type.λi:instr_mode.λop:opcode.λarg:MA_check i.
+ let res ≝ bytes_of_pseudo_instr_mode_param mcu (opcode_to_any mcu op) i arg in
+  λp:defined ? res.
+   let value ≝
+    match res return λres: option (list (t_byte8 mcu)).defined ? res → ? with
+    [ None ⇒ λp:defined (list (t_byte8 mcu)) (None ?).
+       False_rect ? p
+    | Some v ⇒ λ_:defined ? (Some ? v).v
+    ] p in value.
+
+(* detipatore del compilato: (t_byte8 m) → byte8 *)
+definition source_to_byte8 ≝
+λmcu:mcu_type.
+ match mcu
+  return λmcu:mcu_type.list (t_byte8 mcu) → list byte8 with
+  [ HC05 ⇒ λl:list (t_byte8 HC05).
+   let rec aux (p1:list (t_byte8 HC05)) (p2:list byte8) ≝ match p1 with
+    [ nil ⇒ p2 | cons hd tl ⇒ match hd with [ TByte b ⇒ aux tl (p2@[b]) ]] in aux l ([])
+  | HC08 ⇒ λl:list (t_byte8 HC08).
+   let rec aux (p1:list (t_byte8 HC08)) (p2:list byte8) ≝ match p1 with
+    [ nil ⇒ p2 | cons hd tl ⇒ match hd with [ TByte b ⇒ aux tl (p2@[b]) ]] in aux l ([])
+  | HCS08 ⇒ λl:list (t_byte8 HCS08).
+   let rec aux (p1:list (t_byte8 HCS08)) (p2:list byte8) ≝ match p1 with
+    [ nil ⇒ p2 | cons hd tl ⇒ match hd with [ TByte b ⇒ aux tl (p2@[b]) ]] in aux l ([])
+  | RS08 ⇒ λl:list (t_byte8 RS08).
+   let rec aux (p1:list (t_byte8 RS08)) (p2:list byte8) ≝ match p1 with
+    [ nil ⇒ p2 | cons hd tl ⇒ match hd with [ TByte b ⇒ aux tl (p2@[b]) ]] in aux l ([])
+  ].
+
+(* esempio da riciclare su come scrivere un sorgente *)
+(*
+definition source_example_of_correct_HC08: list byte8 ≝
+let m ≝ HC08 in
+ source_to_byte8 m (
+  (compile m ? CLR maINHA I) @
+  (compile m ? NOP maINH I) @
+  (compile m ? BRSETn (maDIRn_and_IMM1 x1 〈x1,x2〉 〈x3,x4〉) I) @
+  (compile m ? ADD maIX0 I) @
+  (compile m ? ADD (maIX1 〈x1,x2〉) I) @
+  (compile m ? ADD (maSP2 〈〈x1,x2〉:〈x3,x4〉〉) I)
+ ).
+*)
diff --git a/helm/software/matita/library/freescale/word16.ma b/helm/software/matita/library/freescale/word16.ma
new file mode 100644 (file)
index 0000000..9a260b1
--- /dev/null
@@ -0,0 +1,402 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+set "baseuri" "cic:/matita/freescale/word16".
+
+(*include "/media/VIRTUOSO/freescale/byte8.ma".*)
+include "freescale/byte8.ma".
+
+(* ********************** *)
+(* DEFINIZIONE DELLE WORD *)
+(* ********************** *)
+
+record word16 : Type ≝
+ {
+ w16h: byte8;
+ w16l: byte8
+ }.
+
+(* \langle \rangle *)
+notation "〈x:y〉" non associative with precedence 80
+ for @{ 'mk_word16 $x $y }.
+interpretation "mk_word16" 'mk_word16 x y = 
+ (cic:/matita/freescale/word16/word16.ind#xpointer(1/1/1) x y).
+
+(* operatore = *)
+definition eq_w16 ≝ λw1,w2.(eq_b8 (w16h w1) (w16h w2)) ⊗ (eq_b8 (w16l w1) (w16l w2)).
+
+(* operatore < *)
+definition lt_w16 ≝
+λw1,w2:word16.match lt_b8 (w16h w1) (w16h w2) with
+ [ true ⇒ true
+ | false ⇒ match gt_b8 (w16h w1) (w16h w2) with
+  [ true ⇒ false
+  | false ⇒ lt_b8 (w16l w1) (w16l w2) ]].
+
+(* operatore ≤ *)
+definition le_w16 ≝ λw1,w2:word16.(eq_w16 w1 w2) ⊕ (lt_w16 w1 w2).
+
+(* operatore > *)
+definition gt_w16 ≝ λw1,w2:word16.⊖ (le_w16 w1 w2).
+
+(* operatore ≥ *)
+definition ge_w16 ≝ λw1,w2:word16.⊖ (lt_w16 w1 w2).
+
+(* operatore and *)
+definition and_w16 ≝
+λw1,w2:word16.mk_word16 (and_b8 (w16h w1) (w16h w2)) (and_b8 (w16l w1) (w16l w2)).
+
+(* operatore or *)
+definition or_w16 ≝
+λw1,w2:word16.mk_word16 (or_b8 (w16h w1) (w16h w2)) (or_b8 (w16l w1) (w16l w2)).
+
+(* operatore xor *)
+definition xor_w16 ≝
+λw1,w2:word16.mk_word16 (xor_b8 (w16h w1) (w16h w2)) (xor_b8 (w16l w1) (w16l w2)).
+
+(* operatore rotazione destra con carry *)
+definition rcr_w16 ≝
+λw:word16.λc:bool.match rcr_b8 (w16h w) c with
+ [ pairT wh' c' ⇒ match rcr_b8 (w16l w) c' with
+  [ pairT wl' c'' ⇒ pairT ?? (mk_word16 wh' wl') c'' ]]. 
+
+(* operatore shift destro *)
+definition shr_w16 ≝
+λw:word16.match rcr_b8 (w16h w) false with
+ [ pairT wh' c' ⇒ match rcr_b8 (w16l w) c' with
+  [ pairT wl' c'' ⇒ pairT ?? (mk_word16 wh' wl') c'' ]].
+
+(* operatore rotazione destra *)
+definition ror_w16 ≝
+λw:word16.match rcr_b8 (w16h w) false with
+ [ pairT wh' c' ⇒ match rcr_b8 (w16l w) c' with
+  [ pairT wl' c'' ⇒ match c'' with
+   [ true ⇒ mk_word16 (or_b8 (mk_byte8 x8 x0) wh') wl'
+   | false ⇒ mk_word16 wh' wl' ]]].
+
+(* operatore rotazione destra n-volte *)
+let rec ror_w16_n (w:word16) (n:nat) on n ≝
+ match n with
+  [ O ⇒ w
+  | S n' ⇒ ror_w16_n (ror_w16 w) n' ].
+
+(* operatore rotazione sinistra con carry *)
+definition rcl_w16 ≝
+λw:word16.λc:bool.match rcl_b8 (w16l w) c with
+ [ pairT wl' c' ⇒ match rcl_b8 (w16h w) c' with
+  [ pairT wh' c'' ⇒ pairT ?? (mk_word16 wh' wl') c'' ]]. 
+
+(* operatore shift sinistro *)
+definition shl_w16 ≝
+λw:word16.match rcl_b8 (w16l w) false with
+ [ pairT wl' c' ⇒ match rcl_b8 (w16h w) c' with
+  [ pairT wh' c'' ⇒ pairT ?? (mk_word16 wh' wl') c'' ]].
+
+(* operatore rotazione sinistra *)
+definition rol_w16 ≝
+λw:word16.match rcl_b8 (w16l w) false with
+ [ pairT wl' c' ⇒ match rcl_b8 (w16h w) c' with
+  [ pairT wh' c'' ⇒ match c'' with
+   [ true ⇒ mk_word16 wh' (or_b8 (mk_byte8 x0 x1) wl')
+   | false ⇒ mk_word16 wh' wl' ]]].
+
+(* operatore rotazione sinistra n-volte *)
+let rec rol_w16_n (w:word16) (n:nat) on n ≝
+ match n with
+  [ O ⇒ w
+  | S n' ⇒ rol_w16_n (rol_w16 w) n' ].
+
+(* operatore not/complemento a 1 *)
+definition not_w16 ≝
+λw:word16.mk_word16 (not_b8 (w16h w)) (not_b8 (w16l w)).
+
+(* operatore somma con carry *)
+definition plus_w16 ≝
+λw1,w2:word16.λc:bool.
+ match plus_b8 (w16l w1) (w16l w2) c with
+  [ pairT l c' ⇒ match plus_b8 (w16h w1) (w16h w2) c' with
+   [ pairT h c'' ⇒ pairT ?? (mk_word16 h l) c'' ]].
+
+(* operatore somma senza carry *)
+definition plus_w16nc ≝
+λw1,w2:word16.fstT ?? (plus_w16 w1 w2 false).
+
+(* operatore carry della somma *)
+definition plus_w16c ≝
+λw1,w2:word16.sndT ?? (plus_w16 w1 w2 false).
+
+(* operatore Most Significant Bit *)
+definition MSB_w16 ≝ λw:word16.eq_ex x8 (and_ex x8 (b8h (w16h w))).
+
+(* word → naturali *)
+definition nat_of_word16 ≝ λw:word16. 256 * (w16h w) + (nat_of_byte8 (w16l w)).
+
+coercion cic:/matita/freescale/word16/nat_of_word16.con.
+
+(* naturali → word *)
+definition word16_of_nat ≝
+ λn.mk_word16 (byte8_of_nat (n / 256)) (byte8_of_nat n).
+
+(* operatore predecessore *)
+definition pred_w16 ≝
+λw:word16.match eq_b8 (w16l w) (mk_byte8 x0 x0) with
+ [ true ⇒ mk_word16 (pred_b8 (w16h w)) (pred_b8 (w16l w))
+ | false ⇒ mk_word16 (w16h w) (pred_b8 (w16l w)) ].
+
+(* operatore successore *)
+definition succ_w16 ≝
+λw:word16.match eq_b8 (w16l w) (mk_byte8 xF xF) with
+ [ true ⇒ mk_word16 (succ_b8 (w16h w)) (succ_b8 (w16l w))
+ | false ⇒ mk_word16 (w16h w) (succ_b8 (w16l w)) ].
+
+(* operatore neg/complemento a 2 *)
+definition compl_w16 ≝
+λw:word16.match MSB_w16 w with
+ [ true ⇒ succ_w16 (not_w16 w)
+ | false ⇒ not_w16 (pred_w16 w) ].
+
+(* 
+   operatore moltiplicazione senza segno: b*b=[0x0000,0xFE01]
+   ... in pratica (〈a,b〉*〈c,d〉) = (a*c)<<8+(a*d)<<4+(b*c)<<4+(b*d)
+*)
+definition mul_b8 ≝
+λb1,b2:byte8.match b1 with
+[ mk_byte8 b1h b1l ⇒ match b2 with
+[ mk_byte8 b2h b2l ⇒ match mul_ex b1l b2l with
+[ mk_byte8 t1_h t1_l ⇒ match mul_ex b1h b2l with
+[ mk_byte8 t2_h t2_l ⇒ match mul_ex b2h b1l with
+[ mk_byte8 t3_h t3_l ⇒ match mul_ex b1h b2h with
+[ mk_byte8 t4_h t4_l ⇒
+ plus_w16nc
+  (plus_w16nc
+   (plus_w16nc 〈〈t4_h,t4_l〉:〈x0,x0〉〉 〈〈x0,t3_h〉:〈t3_l,x0〉〉) 〈〈x0,t2_h〉:〈t2_l,x0〉〉)〈〈x0,x0〉:〈t1_h,t1_l〉〉
+]]]]]].
+
+(* divisione senza segno (secondo la logica delle ALU): (quoziente resto) overflow *)
+definition div_b8 ≝
+λw:word16.λb:byte8.match eq_b8 b 〈x0,x0〉 with
+(* 
+   la combinazione n/0 e' illegale, segnala solo overflow senza dare risultato
+*)
+ [ true ⇒ tripleT ??? 〈xF,xF〉 (w16l w) true
+ | false ⇒ match eq_w16 w 〈〈x0,x0〉:〈x0,x0〉〉 with
+(* 0 diviso qualsiasi cosa diverso da 0 da' q=0 r=0 o=false *)
+  [ true ⇒ tripleT ??? 〈x0,x0〉 〈x0,x0〉 false
+(* 1) e' una divisione sensata che produrra' overflow/risultato *)
+(* 2) parametri: dividendo, divisore, moltiplicatore, quoziente, contatore *)
+(* 3) ad ogni ciclo il divisore e il moltiplicatore vengono scalati di 1 a dx *)
+(* 4) il moltiplicatore e' la quantita' aggiunta al quoziente se il divisore *)
+(*    puo' essere sottratto al dividendo *) 
+  | false ⇒ let rec aux (divd:word16) (divs:word16) (molt:byte8) (q:byte8) (c:nat) on c ≝
+  let w' ≝ plus_w16nc divd (compl_w16 divs) in
+   match c with
+   [ O ⇒ match le_w16 divs divd with
+    [ true ⇒ tripleT ??? (or_b8 molt q) (w16l w') (⊖ (eq_b8 (w16h w') 〈x0,x0〉))
+    | false ⇒ tripleT ??? q (w16l divd) (⊖ (eq_b8 (w16h divd) 〈x0,x0〉)) ]
+   | S c' ⇒ match le_w16 divs divd with
+    [ true ⇒ aux w' (ror_w16 divs) (ror_b8 molt) (or_b8 molt q) c'
+    | false ⇒ aux divd (ror_w16 divs) (ror_b8 molt) q c' ]]
+  in aux w (rol_w16_n 〈〈x0,x0〉:b〉 7) 〈x8,x0〉 〈x0,x0〉 7 ]].
+
+(* operatore x in [inf,sup] *)
+definition in_range ≝
+λx,inf,sup:word16.(le_w16 inf sup) ⊗ (ge_w16 x inf) ⊗ (le_w16 x sup).
+
+(* iteratore sulle word *)
+definition forall_word16 ≝
+ λP.
+  forall_byte8 (λbh.
+  forall_byte8 (λbl.
+   P (mk_word16 bh bl ))).
+
+(* ********************** *)
+(* TEOREMI/LEMMMI/ASSIOMI *)
+(* ********************** *)
+
+(* TODO: dimostrare diversamente *)
+axiom word16_of_nat_nat_of_word16: ∀b. word16_of_nat (nat_of_word16 b) = b.
+
+(* TODO: dimostrare diversamente *)
+axiom lt_nat_of_word16_65536: ∀b. nat_of_word16 b < (256 * 256).
+
+(* TODO: dimostrare diversamente *)
+axiom nat_of_word16_word16_of_nat: ∀n. nat_of_word16 (word16_of_nat n) = n \mod (256 * 256).
+
+(* TODO: dimostrare diversamente *)
+axiom eq_nat_of_word16_n_nat_of_word16_mod_n_65536:
+ ∀n. word16_of_nat n = word16_of_nat (n \mod (256 * 256)).
+
+lemma plusw16_ok:
+ ∀b1,b2,c.
+  match plus_w16 b1 b2 c with
+   [ pairT r c' ⇒ b1 + b2 + nat_of_bool c = nat_of_word16 r + nat_of_bool c' * (256 * 256)
+   ].
+ intros; elim daemon.
+qed.
+
+(* TODO: dimostrare diversamente *)
+axiom plusw16_O_x:
+ ∀b. plus_w16 (mk_word16 (mk_byte8 x0 x0) (mk_byte8 x0 x0)) b false = pairT ?? b false.
+
+lemma plusw16nc_O_x:
+ ∀x. plus_w16nc (mk_word16 (mk_byte8 x0 x0) (mk_byte8 x0 x0)) x = x.
+ intros;
+ unfold plus_w16nc;
+ rewrite > plusw16_O_x;
+ reflexivity.
+qed.
+
+(* TODO: dimostrare diversamente *)
+axiom eq_nat_of_word16_mod: ∀b. nat_of_word16 b = nat_of_word16 b \mod (256 * 256).
+
+(* TODO: dimostrare diversamente *)
+axiom plusw16nc_ok:
+ ∀b1,b2:word16. nat_of_word16 (plus_w16nc b1 b2) = (b1 + b2) \mod (256 * 256).
+
+(* TODO: dimostrare diversamente *)
+axiom eq_eqw16_x0_x0_x0_x0_word16_of_nat_S_false:
+ ∀b. b < (256 * 256 - 1) → eq_w16 (mk_word16 (mk_byte8 x0 x0) (mk_byte8 x0 x0)) (word16_of_nat (S b)) = false.
+
+axiom eq_mod_O_to_exists: ∀n,m. n \mod m = 0 → ∃z. n = z*m.
+
+(* TODO: dimostrare diversamente *)
+axiom eq_w16pred_S_a_a:
+ ∀a. a < (256 * 256 - 1) → pred_w16 (word16_of_nat (S a)) = word16_of_nat a.
+
+(* TODO: dimostrare diversamente *)
+axiom plusw16nc_S:
+ ∀x:word16.∀n.plus_w16nc (word16_of_nat (x*n)) x = word16_of_nat (x * S n).
+
+(* TODO: dimostrare diversamente *)
+axiom eq_plusw16c_x0_x0_x0_x0_x_false:
+ ∀x.plus_w16c (mk_word16 (mk_byte8 x0 x0) (mk_byte8 x0 x0)) x = false.
+
+(* TODO: dimostrare diversamente *)
+axiom eqw16_true_to_eq: ∀b,b'. eq_w16 b b' = true → b=b'.
+
+(* TODO: dimostrare diversamente *)
+axiom eqw16_false_to_not_eq: ∀b,b'. eq_w16 b b' = false → b ≠ b'.
+
+(* TODO: dimostrare diversamente *)
+axiom word16_of_nat_mod: ∀n.word16_of_nat n = word16_of_nat (n \mod (256 * 256)).
+
+(* nuovi *)
+
+(*
+lemma ok_mul_b8: ∀b1,b2:byte8. nat_of_word16 (mul_b8 b1 b2) = b1 * b2.
+ intros;
+ cases b1 (b1h b1l);
+ cases b2 (b2h b2l);
+ change in ⊢ (? ? (? %) ?) with
+  (match mul_ex b1l b2l with
+[ mk_byte8 t1_h t1_l ⇒ match mul_ex b1h b2l with
+[ mk_byte8 t2_h t2_l ⇒ match mul_ex b2h b1l with
+[ mk_byte8 t3_h t3_l ⇒ match mul_ex b1h b2h with
+[ mk_byte8 t4_h t4_l ⇒
+ plus_w16nc
+  (plus_w16nc
+   (plus_w16nc 〈〈t4_h,t4_l〉:〈x0,x0〉〉 〈〈x0,t3_h〉:〈t3_l,x0〉〉) 〈〈x0,t2_h〉:〈t2_l,x0〉〉)〈〈x0,x0〉:〈t1_h,t1_l〉〉
+]]]]);
+ lapply (ok_mul_ex b1l b2l) as ll;
+ lapply (ok_mul_ex b1h b2l) as hl;
+ lapply (ok_mul_ex b2h b1l) as lh;
+ lapply (ok_mul_ex b1h b2h) as hh;
+ elim (mul_ex b1l b2l) (t1_h t1_l);
+ change in ⊢ (? ? (? %) ?) with
+  (match mul_ex b1h b2l with
+[ mk_byte8 t2_h t2_l ⇒ match mul_ex b2h b1l with
+[ mk_byte8 t3_h t3_l ⇒ match mul_ex b1h b2h with
+[ mk_byte8 t4_h t4_l ⇒
+ plus_w16nc
+  (plus_w16nc
+   (plus_w16nc 〈〈t4_h,t4_l〉:〈x0,x0〉〉 〈〈x0,t3_h〉:〈t3_l,x0〉〉) 〈〈x0,t2_h〉:〈t2_l,x0〉〉)〈〈x0,x0〉:〈t1_h,t1_l〉〉
+]]]);
+ elim (mul_ex b1h b2l) (t2_h t2_l);
+ change in ⊢ (? ? (? %) ?) with
+  (match mul_ex b2h b1l with
+[ mk_byte8 t3_h t3_l ⇒ match mul_ex b1h b2h with
+[ mk_byte8 t4_h t4_l ⇒
+ plus_w16nc
+  (plus_w16nc
+   (plus_w16nc 〈〈t4_h,t4_l〉:〈x0,x0〉〉 〈〈x0,t3_h〉:〈t3_l,x0〉〉) 〈〈x0,t2_h〉:〈t2_l,x0〉〉)〈〈x0,x0〉:〈t1_h,t1_l〉〉
+]]);
+ elim (mul_ex b2h b1l) (t3_h t3_l);
+ change in ⊢ (? ? (? %) ?) with
+  (match mul_ex b1h b2h with
+[ mk_byte8 t4_h t4_l ⇒
+ plus_w16nc
+  (plus_w16nc
+   (plus_w16nc 〈〈t4_h,t4_l〉:〈x0,x0〉〉 〈〈x0,t3_h〉:〈t3_l,x0〉〉) 〈〈x0,t2_h〉:〈t2_l,x0〉〉)〈〈x0,x0〉:〈t1_h,t1_l〉〉
+]);
+ elim (mul_ex b1h b2h) (t4_h t4_l);
+ change in ⊢ (? ? (? %) ?) with
+  (plus_w16nc
+  (plus_w16nc
+   (plus_w16nc 〈〈t4_h,t4_l〉:〈x0,x0〉〉 〈〈x0,t3_h〉:〈t3_l,x0〉〉) 〈〈x0,t2_h〉:〈t2_l,x0〉〉)〈〈x0,x0〉:〈t1_h,t1_l〉〉);
+ do 3 (rewrite > plusw16nc_ok);
+ unfold nat_of_word16;
+ unfold nat_of_byte8;
+simplify in ⊢ (? ? (? (? (? (? (? (? (? (? ? (? (? ? (? (? %))) ?)) ?) ?) ?) ?) ?) ?) ?) ?);
+simplify in ⊢ (? ? (? (? (? (? (? (? (? ? (? (? ? (? (? %))) ?)) ?) ?) ?) ?) ?) ?) ?);
+simplify in ⊢ (? ? (? (? (? (? (? (? (? ? (? ? (? (? %)))) ?) ?) ?) ?) ?) ?) ?);
+whd in ⊢ (? ? (? (? (? (? (? (? (? ? (? ? %)) ?) ?) ?) ?) ?) ?) ?);
+whd in ⊢ (? ? (? (? (? (? (? (? (? ? %) ?) ?) ?) ?) ?) ?) ?);
+simplify in ⊢ (? ? (? (? ? (? (? ? (? (? ? (? %)) ?)) ?)) ?) ?);
+simplify in ⊢ (? ? (? (? ? (? (? ? (? ? (? (? %)))) ?)) ?) ?);
+simplify in ⊢ (? ? (? (? ? (? (? ? (? ? %)) ?)) ?) ?);
+whd in ⊢ (? ? (? (? ? (? % ?)) ?) ?);
+simplify in ⊢ (? ? (? (? ? (? ? (? (? ? (? (? %))) ?))) ?) ?);
+simplify in ⊢ (? ? (? (? ? (? ? (? ? (? (? %))))) ?) ?);
+simplify in ⊢ (? ? ? (? (? (? ? (? %)) ?) ?));
+simplify in ⊢ (? ? ? (? (? ? (? %)) ?));
+simplify in ⊢ (? ? ? (? ? (? (? ? (? %)) ?)));
+simplify in ⊢ (? ? ? (? ? (? ? (? %))));
+simplify in ⊢ (? ? (? (? ? (? ? (? (? ? (? %)) ?))) ?) ?);
+simplify in ⊢ (? ? (? (? ? (? ? (? ? (? %)))) ?) ?);
+simplify in ⊢ (? ? (? (? (? (? ? (? ? (? ? (? %)))) ?) ?) ?) ?);
+simplify in ⊢ (? ? (? (? (? (? ? (? ? (? (? ? (? %)) ?))) ?) ?) ?) ?);
+simplify in ⊢ (? ? (? (? (? (? ? (? (? ? (? ? (? %))) ?)) ?) ?) ?) ?);
+simplify in ⊢ (? ? (? (? (? (? ? (? (? ? (? (? ? (? %)) ?)) ?)) ?) ?) ?) ?);
+simplify in ⊢ (? ? (? (? (? (? (? (? ? (? ? (? ? (? %)))) ?) ?) ?) ?) ?) ?);
+simplify in ⊢ (? ? (? (? (? (? (? (? ? (? (? ? (? ? (? %))) ?)) ?) ?) ?) ?) ?) ?);
+simplify in ⊢ (? ? (? (? (? (? (? (? (? (? ? (? (? ? (? %)) ?)) ?) ?) ?) ?) ?) ?) ?) ?);
+simplify in ⊢ (? ? (? (? (? (? (? (? (? (? ? (? ? (? %))) ?) ?) ?) ?) ?) ?) ?) ?);
+simplify in ⊢ (? ? (? (? (? (? (? (? ? (? ? (? (? ? (? %)) ?))) ?) ?) ?) ?) ?) ?);
+simplify in ⊢ (? ? (? (? (? (? (? (? ? (? (? ? (? (? ? (? %)) ?)) ?)) ?) ?) ?) ?) ?) ?);
+rewrite < plus_n_O;
+change in ⊢ (? ? (? (? ? %) ?) ?) with (16*nat_of_exadecim t1_h+nat_of_exadecim t1_l);
+unfold nat_of_byte8 in H H1 H2 H3;
+simplify in ⊢ (? ? (? (? (? (? (? (? ? (? (? ? (? (? ? %) ?)) ?)) ?) ?) ?) ?) ?) ?);
+simplify in ⊢ (? ? (? (? (? (? (? (? ? (? ? (? ? %))) ?) ?) ?) ?) ?) ?);
+simplify in ⊢ (? ? (? (? (? (? ? (? (? ? (? (? ? %) ?)) ?)) ?) ?) ?) ?);
+simplify in ⊢ (? ? (? (? (? (? ? (? ? (? ? %))) ?) ?) ?) ?);
+rewrite < plus_n_O;
+rewrite < plus_n_O;
+simplify in ⊢ (? ? (? (? (? (? (? (? ? (? (? ? %) ?)) ?) ?) ?) ?) ?) ?);
+simplify in ⊢ (? ? (? (? (? (? ? (? (? ? %) ?)) ?) ?) ?) ?);
+elim daemon.
+qed.
+*)