diff --git a/thys/Complx/ex/SumArr.thy b/thys/Complx/ex/SumArr.thy --- a/thys/Complx/ex/SumArr.thy +++ b/thys/Complx/ex/SumArr.thy @@ -1,448 +1,448 @@ (* * Copyright 2016, Data61, CSIRO * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * @TAG(DATA61_BSD) *) section \Case-study\ theory SumArr imports "../OG_Syntax" Word_Lib.Word_Lemmas_32 begin type_synonym routine = nat type_synonym word32 = "32 word" type_synonym funcs = "string \ nat" datatype faults = Overflow | InvalidMem type_synonym 'a array = "'a list" text \Sumarr computes the combined sum of all the elements of multiple arrays. It does this by running a number of threads in parallel, each computing the sum of elements of one of the arrays, and then adding the result to a global variable gsum shared by all threads. \ record sumarr_state = \ \local variables of threads\ tarr :: "routine \ word32 array" tid :: "routine \ word32" ti :: "routine \ word32" tsum :: "routine \ word32" \ \global variables\ glock :: nat gsum :: word32 gdone :: word32 garr :: "(word32 array) array" \ \ghost variables\ ghost_lock :: "routine \ bool" definition NSUM :: word32 where "NSUM = 10" definition MAXSUM :: word32 where "MAXSUM = 1500" definition array_length :: "'a array \ word32" where "array_length arr \ of_nat (length arr)" definition array_nth :: "'a array \ word32 \'a" where "array_nth arr n \ arr ! unat n" definition array_in_bound :: "'a array \ word32 \ bool" where "array_in_bound arr idx \ unat idx < (length arr)" definition - array_nat_sum :: "('a :: len0) word array \ nat" + array_nat_sum :: "('a :: len) word array \ nat" where "array_nat_sum arr \ sum_list (map unat arr)" definition "local_sum arr \ of_nat (min (unat MAXSUM) (array_nat_sum arr))" definition "global_sum arr \ sum_list (map local_sum arr)" definition "tarr_inv s i \ length (tarr s i) = unat NSUM \ tarr s i = garr s ! i" abbreviation "sumarr_inv_till_lock s i \ \gdone s !! i \ ((\ (ghost_lock s) (1 - i)) \ ((gdone s = 0 \ gsum s = 0) \ (gdone s !! (1 - i) \ gsum s = local_sum (garr s !(1 - i)))))" abbreviation "lock_inv s \ (glock s = fromEnum (ghost_lock s 0) + fromEnum (ghost_lock s 1)) \ (\(ghost_lock s) 0 \ \(ghost_lock s) 1)" abbreviation "garr_inv s i \ (\a b. garr s = [a, b]) \ length (garr s ! (1-i)) = unat NSUM" abbreviation "sumarr_inv s i \ lock_inv s \ tarr_inv s i \ garr_inv s i \ tid s i = (of_nat i + 1)" definition lock :: "routine \ (sumarr_state, funcs, faults) ann_com" where "lock i \ \ \sumarr_inv i \ \tsum i = local_sum (\tarr i) \ \sumarr_inv_till_lock i\ AWAIT \glock = 0 THEN \glock:=1,, \ghost_lock:=\ghost_lock (i:= True) END" definition "sumarr_in_lock1 s i \ \gdone s !! i \ ((gdone s = 0 \ gsum s = local_sum (tarr s i)) \ (gdone s !! (1 - i) \ \ gdone s !! i \ gsum s = global_sum (garr s)))" definition "sumarr_in_lock2 s i \ (gdone s !! i \ \ gdone s !! (1 - i) \ gsum s = local_sum (tarr s i)) \ (gdone s !! i \ gdone s !! (1 - i) \ gsum s = global_sum (garr s))" definition unlock :: "routine \ (sumarr_state, funcs, faults) ann_com" where "unlock i \ \ \sumarr_inv i \ \tsum i = local_sum (\tarr i) \ \glock = 1 \ \ghost_lock i \ \gdone !! (unat (\tid i - 1)) \ \sumarr_in_lock2 i \ \\glock := 0,, \ghost_lock:=\ghost_lock (i:= False)\" definition "local_postcond s i \ (\ (ghost_lock s) (1 - i) \ gsum s = (if gdone s !! 0 \ gdone s !! 1 then global_sum (garr s) else local_sum (garr s ! i))) \ gdone s !! i \ \ghost_lock s i" definition sumarr :: "routine \ (sumarr_state, funcs, faults) ann_com" where "sumarr i \ \\sumarr_inv i \ \sumarr_inv_till_lock i\ \tsum:=\tsum(i:=0) ;; \ \tsum i = 0 \ \sumarr_inv i \ \sumarr_inv_till_lock i\ \ti:=\ti(i:=0) ;; TRY \ \tsum i = 0 \ \sumarr_inv i \ \ti i = 0 \ \sumarr_inv_till_lock i\ WHILE \ti i < NSUM INV \ \sumarr_inv i \ \ti i \ NSUM \ \tsum i \ MAXSUM \ \tsum i = local_sum (take (unat (\ti i)) (\tarr i)) \ \sumarr_inv_till_lock i\ DO \ \sumarr_inv i \ \ti i < NSUM \ \tsum i \ MAXSUM \ \tsum i = local_sum (take (unat (\ti i)) (\tarr i)) \ \sumarr_inv_till_lock i\ (InvalidMem, \ array_in_bound (\tarr i) (\ti i) \) \ \ \sumarr_inv i \ \ti i < NSUM \ \tsum i \ MAXSUM \ \tsum i = local_sum (take (unat (\ti i)) (\tarr i)) \ \sumarr_inv_till_lock i\ \tsum:=\tsum(i:=\tsum i + array_nth (\tarr i) (\ti i));; \ \sumarr_inv i \ \ti i < NSUM \ local_sum (take (unat (\ti i)) (\tarr i)) \ MAXSUM \ (\tsum i < MAXSUM \ array_nth (\tarr i) (\ti i) < MAXSUM \ \tsum i = local_sum (take (Suc (unat (\ti i))) (\tarr i))) \ (array_nth (\tarr i) (\ti i) \ MAXSUM \ \tsum i \ MAXSUM\ local_sum (\tarr i) = MAXSUM) \ \sumarr_inv_till_lock i \ (InvalidMem, \ array_in_bound (\tarr i) (\ti i) \) \ \ \sumarr_inv i \ \ti i < NSUM \ (\tsum i < MAXSUM \ array_nth (\tarr i) (\ti i) < MAXSUM \ \tsum i = local_sum (take (Suc (unat (\ti i))) (\tarr i))) \ (array_nth (\tarr i) (\ti i) \ MAXSUM \ \tsum i \ MAXSUM \ local_sum (\tarr i) = MAXSUM) \ \sumarr_inv_till_lock i\ IF array_nth (\tarr i) (\ti i) \ MAXSUM \ \tsum i \ MAXSUM THEN \ \sumarr_inv i \ \ti i < NSUM \ local_sum (\tarr i) = MAXSUM \ \sumarr_inv_till_lock i\ \tsum:=\tsum(i:=MAXSUM);; \ \sumarr_inv i \ \ti i < NSUM \ \tsum i \ MAXSUM \ \tsum i = local_sum (\tarr i) \ \sumarr_inv_till_lock i \ THROW ELSE \ \sumarr_inv i \ \ti i < NSUM \ \tsum i \ MAXSUM \ \tsum i = local_sum (take (Suc (unat (\ti i))) (\tarr i)) \ \sumarr_inv_till_lock i\ SKIP FI;; \ \sumarr_inv i \ \ti i < NSUM \ \tsum i \ MAXSUM \ \tsum i = local_sum (take (Suc (unat (\ti i))) (\tarr i)) \ \sumarr_inv_till_lock i \ \ti:=\ti(i:=\ti i + 1) OD CATCH \ \sumarr_inv i \ \tsum i = local_sum (\tarr i) \ \sumarr_inv_till_lock i\ SKIP END;; \ \sumarr_inv i \ \tsum i = local_sum (\tarr i) \ \sumarr_inv_till_lock i\ SCALL (''lock'', i) 0;; \ \sumarr_inv i \ \tsum i = local_sum (\tarr i) \ \glock = 1 \ \ghost_lock i \ \sumarr_inv_till_lock i \ \gsum:=\gsum + \tsum i ;; \ \sumarr_inv i \ \tsum i = local_sum (\tarr i) \ \glock = 1 \ \ghost_lock i \ \sumarr_in_lock1 i \ \gdone:=(\gdone OR \tid i) ;; \ \sumarr_inv i \ \tsum i = local_sum (\tarr i) \ \glock = 1 \ \ghost_lock i \ \gdone !! (unat (\tid i - 1)) \ \sumarr_in_lock2 i \ SCALL (''unlock'', i) 0" definition precond where "precond s \ (glock s) = 0 \ (gsum s) = 0\ (gdone s) = 0 \ (\a b. garr s = [a, b]) \ (\xs\set (garr s). length xs = unat NSUM) \ (ghost_lock s) 0 = False \ (ghost_lock s) 1 = False" definition postcond where "postcond s \ (gsum s) = global_sum (garr s) \ (\i < 2. (gdone s) !! i)" definition "call_sumarr i \ \length (\garr ! i) = unat NSUM \ \lock_inv \ \garr_inv i \ \sumarr_inv_till_lock i\ CALLX (\s. s\tarr:=(tarr s)(i:=garr s ! i), tid:=(tid s)(i:=of_nat i+1), ti:=(ti s)(i:=undefined), tsum:=(tsum s)(i:=undefined)\) \\sumarr_inv i \ \sumarr_inv_till_lock i\ (''sumarr'', i) 0 (\s t. t\tarr:= (tarr t)(i:=(tarr s) i), tid:=(tid t)(i:=(tid s i)), ti:=(ti t)(i:=(ti s i)), tsum:=(tsum t)(i:=(tsum s i))\) (\_ _. Skip) \\local_postcond i\ \\local_postcond i\ \False\ \False\" definition "\ \ map_of (map (\i. ((''sumarr'', i), com (sumarr i))) [0..<2]) ++ map_of (map (\i. ((''lock'', i), com (lock i))) [0..<2]) ++ map_of (map (\i. ((''unlock'', i), com (unlock i))) [0..<2])" definition "\ \ map_of (map (\i. ((''sumarr'', i), [ann (sumarr i)])) [0..<2]) ++ map_of (map (\i. ((''lock'', i), [ann (lock i)])) [0..<2]) ++ map_of (map (\i. ((''unlock'', i), [ann (unlock i)])) [0..<2])" declare [[goals_limit = 10]] lemma [simp]: "local_sum [] = 0" by (simp add: local_sum_def array_nat_sum_def) lemma MAXSUM_le_plus: "x < MAXSUM \ MAXSUM \ MAXSUM + x" unfolding MAXSUM_def apply (rule word_le_plus[rotated], assumption) apply clarsimp done lemma local_sum_Suc: "\n < length arr; local_sum (take n arr) + arr ! n < MAXSUM; arr ! n < MAXSUM\ \ local_sum (take n arr) + arr ! n = local_sum (take (Suc n) arr)" apply (subst take_Suc_conv_app_nth) apply clarsimp apply (clarsimp simp: local_sum_def array_nat_sum_def ) apply (subst (asm) min_def, clarsimp split: if_splits) apply (clarsimp simp: MAXSUM_le_plus word_not_le[symmetric]) apply (subst min_absorb2) apply (subst of_nat_mono_maybe_le[where 'a=32]) apply (clarsimp simp: MAXSUM_def) apply (clarsimp simp: MAXSUM_def) apply unat_arith apply (clarsimp simp: MAXSUM_def) apply unat_arith apply clarsimp done lemma local_sum_MAXSUM: "k < length arr \ MAXSUM \ arr ! k \ local_sum arr = MAXSUM" apply (clarsimp simp: local_sum_def array_nat_sum_def) apply (rule word_unat.Rep_inverse') apply (rule min_absorb1[symmetric]) apply (subst (asm) word_le_nat_alt) apply (rule le_trans[rotated]) apply (rule elem_le_sum_list) apply simp apply clarsimp done lemma local_sum_MAXSUM': "k < length arr \ MAXSUM \ local_sum (take k arr) + arr ! k \ local_sum (take k arr) \ MAXSUM \ arr ! k \ MAXSUM \ local_sum arr = MAXSUM" apply (clarsimp simp: local_sum_def array_nat_sum_def) apply (rule word_unat.Rep_inverse') apply (rule min_absorb1[symmetric]) apply (subst (asm) word_le_nat_alt) apply (subst (asm) unat_plus_simple[THEN iffD1]) apply (rule word_add_le_mono2[where i=0, simplified]) apply (clarsimp simp: MAXSUM_def) apply unat_arith apply (rule le_trans, assumption) apply (subst unat_of_nat_eq) apply (clarsimp simp: MAXSUM_def min.strict_coboundedI1) by (subst id_take_nth_drop[where i=k], simp, clarsimp simp: trans_le_add1) lemma word_min_0[simp]: - "min (x::'a::len0 word) 0 = 0" - "min 0 (x::'a::len0 word) = 0" + "min (x::'a::len word) 0 = 0" + "min 0 (x::'a::len word) = 0" by (simp add:min_def)+ ML \fun TRY' tac i = TRY (tac i)\ lemma imp_disjL_context': "((P \ R) \ (Q \ R)) = ((P \ R) \ (\P \ Q \ R))" by auto lemma map_of_prod_1[simp]: "i < n \ map_of (map (\i. ((p, i), g i)) [0.. p \ q \ (m ++ map_of (map (\i. ((p, i), g i)) [0.. \ (''sumarr'',n) = Some (com (sumarr n))" "n < 2 \ \ (''sumarr'',n) = Some ([ann (sumarr n)])" "n < 2 \ \ (''lock'',n) = Some (com (lock n))" "n < 2 \ \ (''lock'',n) = Some ([ann (lock n)])" "n < 2 \ \ (''unlock'',n) = Some (com (unlock n))" "n < 2 \ \ (''unlock'',n) = Some ([ann (unlock n)])" "[ann (sumarr n)]!0 = ann (sumarr n)" "[ann (lock n)]!0 = ann (lock n)" "[ann (unlock n)]!0 = ann (unlock n)" by (simp add: \_def \_def)+ lemmas sumarr_proc_simp_unfolded = sumarr_proc_simp[unfolded sumarr_def unlock_def lock_def oghoare_simps] lemma oghoare_sumarr: notes sumarr_proc_simp_unfolded[proc_simp add] shows "i < 2 \ \, \ |\\<^bsub>/F\<^esub> sumarr i \\local_postcond i\, \False\" unfolding sumarr_def unlock_def lock_def ann_call_def call_def block_def apply simp apply oghoare (*23*) unfolding tarr_inv_def array_length_def array_nth_def array_in_bound_def sumarr_in_lock1_def sumarr_in_lock2_def apply (tactic "PARALLEL_ALLGOALS ((TRY' o SOLVED') (clarsimp_tac (@{context} addsimps @{thms local_postcond_def global_sum_def ex_in_conv[symmetric]}) THEN_ALL_NEW fast_force_tac (@{context} addSDs @{thms less_2_cases} addIs @{thms local_sum_Suc unat_mono} ) ))") (*2*) apply clarsimp apply (rule conjI) apply (fastforce intro!: local_sum_Suc unat_mono) apply (subst imp_disjL_context') apply (rule conjI) apply clarsimp apply (erule local_sum_MAXSUM[rotated]) apply unat_arith apply (clarsimp simp: not_le) apply (erule (1) local_sum_MAXSUM'[rotated] ; unat_arith) apply clarsimp apply unat_arith done lemma less_than_two_2[simp]: "i < 2 \ Suc 0 - i < 2" by arith lemma oghoare_call_sumarr: notes sumarr_proc_simp[proc_simp add] shows "i < 2 \ \, \ |\\<^bsub>/F\<^esub> call_sumarr i \\local_postcond i\, \False\" unfolding call_sumarr_def ann_call_def call_def block_def tarr_inv_def apply oghoare (*10*) apply (clarsimp; fail | ((simp only: pre.simps)?, rule oghoare_sumarr))+ apply (clarsimp simp: sumarr_def tarr_inv_def) apply (clarsimp simp: local_postcond_def; fail)+ done lemma less_than_two_inv[simp]: "i < 2 \ j < 2 \ i \ j \ Suc 0 - i = j" by simp lemma inter_aux_call_sumarr[simplified]: notes sumarr_proc_simp_unfolded[proc_simp add] com.simps[oghoare_simps add] shows "i < 2 \ j < 2 \ i \ j \ interfree_aux \ \ F (com (call_sumarr i), (ann (call_sumarr i), \\local_postcond i\, \False\), com (call_sumarr j), ann (call_sumarr j))" unfolding call_sumarr_def ann_call_def call_def block_def tarr_inv_def sumarr_def lock_def unlock_def apply oghoare_interfree_aux (*650*) unfolding tarr_inv_def local_postcond_def sumarr_in_lock1_def sumarr_in_lock2_def by (tactic "PARALLEL_ALLGOALS ( TRY' (remove_single_Bound_mem @{context}) THEN' (TRY' o SOLVED') (clarsimp_tac @{context} THEN_ALL_NEW fast_force_tac (@{context} addSDs @{thms less_2_cases}) ))") (* 2 minutes *) lemma pre_call_sumarr: "i < 2 \ precond x \ x \ pre (ann (call_sumarr i))" unfolding precond_def call_sumarr_def ann_call_def by (fastforce dest: less_2_cases simp: array_length_def) lemma post_call_sumarr: "local_postcond x 0 \ local_postcond x 1 \ postcond x" unfolding postcond_def local_postcond_def by (fastforce dest: less_2_cases split: if_splits) lemma sumarr_correct: "\, \ |\\<^bsub>/F\<^esub> \\precond\ COBEGIN SCHEME [0 \ m < 2] call_sumarr m \\local_postcond m\,\False\ COEND \\postcond\, \False\" apply oghoare (* 5 subgoals *) apply (fastforce simp: pre_call_sumarr) apply (rule oghoare_call_sumarr, simp) apply (clarsimp simp: post_call_sumarr) apply (simp add: inter_aux_call_sumarr) apply clarsimp done end diff --git a/thys/Containers/ITP-2013/Benchmark_Set.thy b/thys/Containers/ITP-2013/Benchmark_Set.thy --- a/thys/Containers/ITP-2013/Benchmark_Set.thy +++ b/thys/Containers/ITP-2013/Benchmark_Set.thy @@ -1,80 +1,80 @@ theory Benchmark_Set imports "HOL-Word.Word" "HOL-Library.Cardinality" begin instantiation word :: (len) card_UNIV begin definition "finite_UNIV = Phantom('a word) True" definition "card_UNIV = Phantom('a word) (2 ^ LENGTH('a))" instance by(intro_classes)(simp_all add: card_UNIV_word_def card_word finite_UNIV_word_def) end -definition word_of_integer :: "integer \ 'a::len0 word" +definition word_of_integer :: "integer \ 'a::len word" where "word_of_integer = word_of_int \ int_of_integer" lemma word_of_integer_code [code]: "word_of_integer k = (if k < 0 then - word_of_integer (- k) else if k = 0 then 0 else let (q, r) = divmod_integer k 2 in if r = 0 then 2 * word_of_integer q else 2 * word_of_integer q + 1)" apply(unfold word_of_integer_def o_def) apply(subst int_of_integer_code) apply(clarsimp simp add: divmod_integer_def) by (metis minus_minus plus_word.abs_eq times_word.abs_eq wi_hom_neg word_1_wi word_numeral_alt) -definition word_of :: "natural \ 'a::len0 word" +definition word_of :: "natural \ 'a::len word" where "word_of = word_of_integer o integer_of_natural" text \randomly generate a set of (up to) n elements drawn from 0 to bound\ fun gen_build1 :: "natural \ nat \ (32 word set \ Random.seed) \ (32 word set \ Random.seed)" where "gen_build1 bound n (A, seed) = (if n = 0 then (A, seed) else let (x, seed') = Random.range bound seed in gen_build1 bound (n - 1) (insert (word_of x) A, seed'))" declare gen_build1.simps[simp del] definition build1 :: "natural \ Random.seed \ (32 word set \ Random.seed)" where "build1 bound seed = (let (n', seed') = Random.range bound seed; (compl, seed'') = Random.range 2 seed; (x, seed''') = gen_build1 bound (Code_Numeral.nat_of_natural n') ({}, seed'') in (if compl = 0 then x else - x, seed'''))" text \randomly generate a set of (up to) n sets each with a random number between 0 and bound of elements between 0 and bound\ fun gen_build2 :: "natural \ nat \ (32 word set set \ Random.seed) \ (32 word set set \ Random.seed)" where "gen_build2 bound n (A, seed) = (if n = 0 then (A, seed) else let (x, seed') = build1 bound seed in gen_build2 bound (n - 1) (insert x A, seed'))" declare gen_build2.simps[simp del] definition build :: "nat \ nat \ Random.seed \ 32 word set set \ Random.seed" where "build n m seed = gen_build2 (of_nat m) n ({}, seed)" fun gen_lookup :: "32 word set set \ natural \ nat \ (nat \ Random.seed) \ (nat \ Random.seed)" where "gen_lookup A bound n (hits, seed) = (if n = 0 then (hits, seed) else let (x, seed') = build1 bound seed in gen_lookup A bound (n - 1) (if x : A then hits + 1 else hits, seed'))" declare gen_lookup.simps [simp del] primrec lookup :: "nat \ nat \ (32 word set set \ Random.seed) \ (nat \ Random.seed)" where "lookup n m (A, seed) = gen_lookup A (of_nat m) 100 (0, seed)" definition covered :: "32 word set set \ nat" where "covered A = card (\A)" definition complete :: "nat \ nat \ Random.seed \ (nat \ nat)" where "complete n m seed = (let (A, seed') = build n m seed in (fst (lookup n m (A, seed)), covered A))" end diff --git a/thys/Containers/ITP-2013/Benchmark_Set_LC.thy b/thys/Containers/ITP-2013/Benchmark_Set_LC.thy --- a/thys/Containers/ITP-2013/Benchmark_Set_LC.thy +++ b/thys/Containers/ITP-2013/Benchmark_Set_LC.thy @@ -1,88 +1,88 @@ theory Benchmark_Set_LC imports Benchmark_Set Containers.Set_Impl "HOL-Library.Code_Target_Nat" begin lemma [code_unfold del]: "card \ Cardinality.card'" by(simp) -instantiation word :: (len0) ceq begin +instantiation word :: (len) ceq begin definition "CEQ('a word) = Some (=)" instance by(intro_classes)(simp add: ceq_word_def) end -instantiation word :: (len0) compare begin +instantiation word :: (len) compare begin definition "compare_word = (comparator_of :: 'a word comparator)" instance by(intro_classes)(simp add: compare_word_def comparator_of) end -instantiation word :: (len0) ccompare begin +instantiation word :: (len) ccompare begin definition "CCOMPARE('a word) = Some compare" instance by(intro_classes)(simp add: ccompare_word_def comparator_compare) end -instantiation word :: (len0) set_impl begin +instantiation word :: (len) set_impl begin definition "SET_IMPL('a word) = Phantom('a word) set_RBT" instance .. end instantiation word :: (len) proper_interval begin fun proper_interval_word :: "'a word option \ 'a word option \ bool" where "proper_interval_word None None = True" | "proper_interval_word None (Some y) = (y \ 0)" | "proper_interval_word (Some x) None = (x \ max_word)" | "proper_interval_word (Some x) (Some y) = (x < y \ x \ y - 1)" instance proof let ?pi = "proper_interval :: 'a word proper_interval" show "?pi None None = True" by simp fix y show "?pi None (Some y) = (\z. z < y)" by simp (metis word_gt_0 word_not_simps(1)) fix x show "?pi (Some x) None = (\z. x < z)" by simp (metis eq_iff max_word_max not_le) have "(x < y \ x \ y - 1) = (\z>x. z < y)" (is "?lhs \ ?rhs") proof assume ?lhs then obtain "x < y" "x \ y - 1" .. have "0 \ uint x" by simp also have "\ < uint y" using \x < y\ by(simp add: word_less_def) finally have "0 < uint y" . then have "y - 1 < y" by(simp add: word_less_def uint_sub_if' not_le) moreover from \0 < uint y\ \x < y\ \x \ y - 1\ have "x < y - 1" by(simp add: word_less_def uint_sub_if' uint_arith_simps(3)) ultimately show ?rhs by blast next assume ?rhs then obtain z where z: "x < z" "z < y" by blast have "0 \ uint z" by simp also have "\ < uint y" using \z < y\ by(simp add: word_less_def) finally show ?lhs using z by(auto simp add: word_less_def uint_sub_if') qed thus "?pi (Some x) (Some y) = (\z>x. z < y)" by simp qed end instantiation word :: (len) cproper_interval begin definition "cproper_interval = (proper_interval :: 'a word proper_interval)" instance by( intro_classes, simp add: cproper_interval_word_def ccompare_word_def compare_word_def le_lt_comparator_of ID_Some proper_interval_class.axioms) end -instantiation word :: (len0) cenum begin +instantiation word :: (len) cenum begin definition "CENUM('a word) = None" instance by(intro_classes)(simp_all add: cEnum_word_def) end notepad begin have "Benchmark_Set.complete 30 40 (12345, 67899) = (32, 4294967296)" by eval end end diff --git a/thys/IEEE_Floating_Point/Conversion_IEEE_Float.thy b/thys/IEEE_Floating_Point/Conversion_IEEE_Float.thy --- a/thys/IEEE_Floating_Point/Conversion_IEEE_Float.thy +++ b/thys/IEEE_Floating_Point/Conversion_IEEE_Float.thy @@ -1,346 +1,346 @@ (* Author: Fabian Hellauer Fabian Immler *) theory Conversion_IEEE_Float imports "HOL-Library.Float" IEEE_Properties "HOL-Library.Code_Target_Numeral" begin definition "of_finite (x::('e, 'f)float) = (if is_normal x then (Float (normal_mantissa x) (normal_exponent x)) else if is_denormal x then (Float (denormal_mantissa x) (denormal_exponent TYPE(('e, 'f)float))) else 0)" lemma float_val_of_finite: "is_finite x \ of_finite x = valof x" by (induction x) (auto simp: normal_imp_not_denormal of_finite_def) definition is_normal_Float::"('e, 'f)float itself \ Float.float \ bool" where "is_normal_Float x f \ mantissa f \ 0 \ bitlen \mantissa f\ \ fracwidth x + 1 \ - int (bias x) - bitlen \mantissa f\ + 1 < Float.exponent f \ Float.exponent f < 2^(LENGTH('e)) - bitlen \mantissa f\ - bias x" definition is_denormal_Float::"('e, 'f)float itself \ Float.float \ bool" where "is_denormal_Float x f \ mantissa f \ 0 \ bitlen \mantissa f\ \ 1 - Float.exponent f - int (bias x) \ 1 - 2^(LENGTH('e) - 1) - int LENGTH('f) < Float.exponent f" lemmas is_denormal_FloatD = is_denormal_Float_def[THEN iffD1, THEN conjunct1] is_denormal_Float_def[THEN iffD1, THEN conjunct2] definition is_finite_Float::"('e, 'f)float itself \ Float.float \ bool" where "is_finite_Float x f \ is_normal_Float x f \ is_denormal_Float x f \ f = 0" lemma is_finite_Float_eq: "is_finite_Float TYPE(('e, 'f)float) f \ (let e = Float.exponent f; bm = bitlen (abs (mantissa f)) in bm \ Suc LENGTH('f) \ bm \ 2 ^ (LENGTH('e) - 1) - e \ 1 - 2 ^ (LENGTH('e) - 1) - int LENGTH('f) < e)" proof - have *: "(2::int) ^ (LENGTH('e) - Suc 0) - 1 < 2 ^ LENGTH('e)" by (metis Suc_1 diff_le_self lessI linorder_not_less one_less_numeral_iff power_strict_increasing_iff zle_diff1_eq) have **: "1 - 2 ^ (LENGTH('e) - Suc 0) < int LENGTH('f)" by (smt len_gt_0 of_nat_0_less_iff zero_less_power) have ***: "2 ^ (LENGTH('e) - 1) + 1 = 2 ^ LENGTH('e) - int (bias TYPE(('e, 'f) IEEE.float))" by (simp add: bias_def power_Suc[symmetric]) have rewr: "x \ 2 ^ n - e \ x + e < 2 ^ n + 1" for x::int and n e by auto show ?thesis unfolding *** rewr using * ** unfolding is_finite_Float_def is_normal_Float_def is_denormal_Float_def by (auto simp: Let_def bias_def mantissa_eq_zero_iff intro: le_less_trans[OF add_right_mono]) qed lift_definition normal_of_Float :: "Float.float \ ('e, 'f)float" is "\x. let m = mantissa x; e = Float.exponent x in (if m > 0 then 0 else 1, word_of_int (e + int (bias TYPE(('e, 'f)float)) + bitlen \m\ - 1), word_of_int (\m\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \m\)) - 2 ^ (LENGTH('f))))" . lemma sign_normal_of_Float:"sign (normal_of_Float x) = (if x > 0 then 0 else 1)" by transfer (auto simp: Let_def mantissa_pos_iff) lemma uints_bitlen_eq: "uints n = {i. 0 \ i \ bitlen i \ n}" by (auto simp: uints_num bitlen_le_iff_power) lemma uint_word_of_int_bitlen_eq: - "uint (word_of_int x::'a::len0 word) = x" if "bitlen x \ LENGTH('a)" "x \ 0" + "uint (word_of_int x::'a::len word) = x" if "bitlen x \ LENGTH('a)" "x \ 0" by (subst word_uint.Abs_inverse) (simp_all add: uints_bitlen_eq that) lemma fraction_normal_of_Float:"fraction (normal_of_Float x::('e, 'f)float) = (nat \mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\)) - 2 ^ LENGTH('f))" if "is_normal_Float TYPE(('e, 'f)float) x" proof - from that have bmp: "bitlen \mantissa x\ > 0" by (metis abs_of_nonneg bitlen_bounds bitlen_def is_normal_Float_def nat_code(2) of_nat_0_le_iff power.simps(1) zabs_less_one_iff zero_less_abs_iff) have mless: "\mantissa x\ < 2 ^ nat (bitlen \mantissa x\)" using bitlen_bounds by force have lem: "2 ^ nat (bitlen \mantissa x\ - 1) \ \mantissa x\" using bitlen_bounds is_normal_Float_def that zero_less_abs_iff by blast from that have nble: "nat (bitlen \mantissa x\) \ Suc LENGTH('f)" using bitlen_bounds by (auto simp: is_normal_Float_def) have nn: "0 \ \mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\)) - 2 ^ LENGTH('f)" apply (rule add_le_imp_le_diff) apply (rule order_trans[rotated]) apply (rule mult_right_mono) apply (rule lem, force) unfolding power_add[symmetric] using nble bmp by (auto) have "\mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\)) < 2 * 2 ^ LENGTH('f)" apply (rule less_le_trans) apply (rule mult_strict_right_mono) apply (rule mless) apply force unfolding power_add[symmetric] power_Suc[symmetric] apply (rule power_increasing) using nble by auto then have "bitlen (\mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\)) - 2 ^ LENGTH('f)) \ int LENGTH('f)" unfolding bitlen_le_iff_power by simp then show ?thesis apply (transfer fixing: x) unfolding Let_def split_beta' fst_conv snd_conv uint_nat[symmetric] unat_def using nn by (subst uint_word_of_int_bitlen_eq) (auto simp: nat_mult_distrib nat_diff_distrib nat_power_eq) qed lemma exponent_normal_of_Float:"exponent (normal_of_Float x::('e, 'f)float) = nat (Float.exponent x + (bias TYPE(('e, 'f)float)) + bitlen \mantissa x\ - 1)" if "is_normal_Float TYPE(('e, 'f)float) x" using that by (transfer fixing: x) (auto simp: is_normal_Float_def bitlen_le_iff_power uint_word_of_int_bitlen_eq Let_def uint_nat[symmetric] unat_def) lift_definition denormal_of_Float :: "Float.float \ ('e, 'f)float" is "\x. let m = mantissa x; e = Float.exponent x in (if m \ 0 then 0 else 1, 0, word_of_int (\m\ * 2 ^ nat (e + bias TYPE(('e, 'f)float) + fracwidth TYPE(('e, 'f)float) - 1)))" . lemma sign_denormal_of_Float:"sign (denormal_of_Float x) = (if x \ 0 then 0 else 1)" by transfer (auto simp: Let_def mantissa_nonneg_iff) lemma exponent_denormal_of_Float:"exponent (denormal_of_Float x::('e, 'f)float) = 0" by (transfer fixing: x) (auto simp: Let_def) lemma fraction_denormal_of_Float:"fraction (denormal_of_Float x::('e, 'f)float) = (nat \mantissa x\ * 2 ^ nat (Float.exponent x + bias TYPE(('e, 'f)float) + LENGTH('f) - 1))" if "is_denormal_Float TYPE(('e, 'f)float) x" proof - have mless: "\mantissa x\ < 2 ^ nat (bitlen \mantissa x\)" using bitlen_bounds by force have *: "nat (bitlen \mantissa x\) + nat (Float.exponent x + (2 ^ (LENGTH('e) - Suc 0) + int LENGTH('f)) - 2) \ LENGTH('f)" using that by (auto simp: is_denormal_Float_def nat_diff_distrib' le_diff_conv bitlen_nonneg nat_le_iff bias_def nat_add_distrib[symmetric]) have "\mantissa x\ * 2 ^ nat (Float.exponent x + int (bias TYPE(('e, 'f)float)) + LENGTH('f) - 1) < 2 ^ LENGTH('f)" apply (rule less_le_trans) apply (rule mult_strict_right_mono) apply (rule mless, force) unfolding power_add[symmetric] power_Suc[symmetric] apply (rule power_increasing) apply (auto simp: bias_def) using that * by (auto simp: is_denormal_Float_def algebra_simps) then show ?thesis apply (transfer fixing: x) unfolding Let_def split_beta' fst_conv snd_conv uint_nat[symmetric] unat_def apply (subst uint_word_of_int_bitlen_eq) unfolding bitlen_le_iff_power by (auto simp: nat_mult_distrib) qed definition of_finite_Float :: "Float.float \ ('e, 'f) float" where "of_finite_Float x = (if is_normal_Float TYPE(('e, 'f)float) x then normal_of_Float x else if is_denormal_Float TYPE(('e, 'f)float) x then denormal_of_Float x else 0)" lemma valof_normal_of_Float: "valof (normal_of_Float x::('e, 'f)float) = x" if "is_normal_Float TYPE(('e, 'f)float) x" proof - have "valof (normal_of_Float x::('e, 'f)float) = (- 1) ^ sign (normal_of_Float x::('e, 'f)float) * ((1 + real (nat \mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\)) - 2 ^ LENGTH('f)) / 2 ^ LENGTH('f)) * 2 powr (bitlen \mantissa x\ - 1)) * 2 powr Float.exponent x" (is "_ = ?s * ?m * ?e") using that by (auto simp: is_normal_Float_def valof_eq fraction_normal_of_Float powr_realpow[symmetric] exponent_normal_of_Float powr_diff powr_add) also have "\mantissa x\ > 0" using that by (auto simp: is_normal_Float_def) have bound: "2 ^ LENGTH('f) \ nat \mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\))" proof - have "(2::nat) ^ LENGTH('f) \ 2 ^ nat (bitlen \mantissa x\ - 1) * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\))" by (simp add: power_add[symmetric]) also have "\ \ nat \mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\))" using bitlen_bounds[of "\mantissa x\"] that by (auto simp: is_normal_Float_def) finally show ?thesis . qed have "?m = abs (mantissa x)" apply (subst of_nat_diff) subgoal using bound by auto subgoal using that by (auto simp: powr_realpow[symmetric] powr_add[symmetric] is_normal_Float_def bitlen_nonneg of_nat_diff divide_simps) done finally show ?thesis by (auto simp: mantissa_exponent sign_normal_of_Float abs_real_def zero_less_mult_iff) qed lemma valof_denormal_of_Float: "valof (denormal_of_Float x::('e, 'f)float) = x" if "is_denormal_Float TYPE(('e, 'f)float) x" proof - have less: "0 < Float.exponent x + (int (bias TYPE(('e, 'f) IEEE.float)) + int LENGTH('f))" using that by (auto simp: is_denormal_Float_def bias_def) have "valof (denormal_of_Float x::('e, 'f)float) = ((- 1) ^ sign (denormal_of_Float x::('e, 'f)float) * \real_of_int (mantissa x)\) * (2 powr real (nat (Float.exponent x + int (bias TYPE(('e, 'f) IEEE.float)) + int LENGTH('f) - 1)) / (2 powr real (bias TYPE(('e, 'f) IEEE.float)) * 2 powr LENGTH('f)) * 2)" (is "_ = ?m * ?e") by (auto simp: valof_eq exponent_denormal_of_Float fraction_denormal_of_Float that mantissa_exponent powr_realpow[symmetric]) also have "?m = mantissa x" by (auto simp: sign_denormal_of_Float abs_real_def mantissa_neg_iff) also have "?e = 2 powr Float.exponent x" by (auto simp: powr_add[symmetric] divide_simps powr_mult_base less ac_simps) finally show ?thesis by (simp add: mantissa_exponent) qed lemma valof_of_finite_Float: "is_finite_Float (TYPE(('e, 'f) IEEE.float)) x \ valof (of_finite_Float x::('e, 'f)float) = x" by (auto simp: of_finite_Float_def is_finite_Float_def valof_denormal_of_Float valof_normal_of_Float) lemma is_normal_normal_of_Float: "is_normal (normal_of_Float x::('e, 'f)float)" if "is_normal_Float TYPE(('e, 'f)float) x" using that by (auto simp: is_normal_def exponent_normal_of_Float that is_normal_Float_def emax_eq nat_less_iff) lemma is_denormal_denormal_of_Float: "is_denormal (denormal_of_Float x::('e, 'f)float)" if "is_denormal_Float TYPE(('e, 'f)float) x" using that by (auto simp: is_denormal_def exponent_denormal_of_Float that is_denormal_Float_def emax_eq fraction_denormal_of_Float le_nat_iff bias_def) lemma is_finite_of_finite_Float: "is_finite (of_finite_Float x)" by (auto simp: is_finite_def of_finite_Float_def is_normal_normal_of_Float is_denormal_denormal_of_Float) lemma Float_eq_zero_iff: "Float m e = 0 \ m = 0" by (metis Float.compute_is_float_zero Float_0_eq_0) lemma bitlen_mantissa_Float: shows "bitlen \mantissa (Float m e)\ = (if m = 0 then 0 else bitlen \m\ + e) - Float.exponent (Float m e)" using bitlen_Float[of m e] by auto lemma exponent_Float: shows "Float.exponent (Float m e) = (if m = 0 then 0 else bitlen \m\ + e) - bitlen \mantissa (Float m e)\ " using bitlen_Float[of m e] by auto lemma is_normal_Float_normal: "is_normal_Float TYPE(('e, 'f)float) (Float (normal_mantissa x) (normal_exponent x))" if "is_normal x" for x::"('e, 'f)float" proof - define f where "f = Float (normal_mantissa x) (normal_exponent x)" from that have "f \ 0" by (auto simp: f_def is_normal_def zero_float_def[symmetric] Float_eq_zero_iff normal_mantissa_def add_nonneg_eq_0_iff) from denormalize_shift[OF f_def this] obtain i where i: "normal_mantissa x = mantissa f * 2 ^ i" "normal_exponent x = Float.exponent f - int i" by auto have "mantissa f \ 0" by (auto simp: \f \ 0\ i mantissa_eq_zero_iff Float_eq_zero_iff) moreover have "normal_exponent x \ Float.exponent f" unfolding i by simp then have " bitlen \mantissa f\ \ 1 + int LENGTH('f)" unfolding bitlen_mantissa_Float bitlen_normal_mantissa f_def by auto moreover have "- int (bias TYPE(('e, 'f)float)) - bitlen \mantissa f\ + 1 < Float.exponent f" unfolding bitlen_mantissa_Float bitlen_normal_mantissa f_def using that by (auto simp: mantissa_eq_zero_iff abs_mult bias_def normal_mantissa_def normal_exponent_def is_normal_def emax_eq less_diff_conv add_nonneg_eq_0_iff) moreover have "2 ^ (LENGTH('e) - Suc 0) + - (1::int) * 2 ^ LENGTH('e) \ 0" by simp then have "(2::int) ^ (LENGTH('e) - Suc 0) < 1 + 2 ^ LENGTH('e)" by arith then have "Float.exponent f < 2 ^ LENGTH('e) - bitlen \mantissa f\ - int (bias TYPE(('e, 'f)float))" using normal_exponent_bounds_int[OF that] unfolding bitlen_mantissa_Float bitlen_normal_mantissa f_def by (auto simp: bias_def algebra_simps power_Suc[symmetric] intro: le_less_trans[OF add_right_mono] normal_exponent_bounds_int[OF that]) ultimately show ?thesis by (auto simp: is_normal_Float_def f_def) qed lemma is_denormal_Float_denormal: "is_denormal_Float TYPE(('e, 'f)float) (Float (denormal_mantissa x) (denormal_exponent TYPE(('e, 'f)float)))" if "is_denormal x" for x::"('e, 'f)float" proof - define f where "f = Float (denormal_mantissa x) (denormal_exponent TYPE(('e, 'f)float))" from that have "f \ 0" by (auto simp: f_def is_denormal_def zero_float_def[symmetric] Float_eq_zero_iff denormal_mantissa_def add_nonneg_eq_0_iff) from denormalize_shift[OF f_def this] obtain i where i: "denormal_mantissa x = mantissa f * 2 ^ i" "denormal_exponent TYPE(('e, 'f)float) = Float.exponent f - int i" by auto have "mantissa f \ 0" by (auto simp: \f \ 0\ i mantissa_eq_zero_iff Float_eq_zero_iff) moreover have "bitlen \mantissa f\ \ 1 - Float.exponent f - int (bias TYPE(('e, 'f) IEEE.float))" using \mantissa f \ 0\ unfolding f_def bitlen_mantissa_Float using bitlen_denormal_mantissa[of x] by (auto simp: denormal_exponent_def) moreover have "2 - 2 ^ (LENGTH('e) - Suc 0) - int LENGTH('f) \ Float.exponent f" (is "?l \ _") proof - have "?l \ denormal_exponent TYPE(('e, 'f)float) + i" using that by (auto simp: is_denormal_def bias_def denormal_exponent_def) also have "\ = Float.exponent f" unfolding i by auto finally show ?thesis . qed ultimately show ?thesis unfolding is_denormal_Float_def exponent_Float f_def[symmetric] by auto qed lemma is_finite_Float_of_finite: "is_finite_Float TYPE(('e, 'f)float) (of_finite x)" for x::"('e, 'f)float" by (auto simp: is_finite_Float_def of_finite_def is_normal_Float_normal is_denormal_Float_denormal) end diff --git a/thys/IP_Addresses/WordInterval.thy b/thys/IP_Addresses/WordInterval.thy --- a/thys/IP_Addresses/WordInterval.thy +++ b/thys/IP_Addresses/WordInterval.thy @@ -1,775 +1,775 @@ (* Title: WordInterval.thy Authors: Julius Michaelis, Cornelius Diekmann *) theory WordInterval imports Main "Word_Lib.Word_Lemmas" begin section\WordInterval: Executable datatype for Machine Word Sets\ text\Stores ranges of machine words as interval. This has been proven quite efficient for IP Addresses.\ (*NOTE: All algorithms here use a straight-forward implementation. There is a lot of room for improving the computation complexity, for example by making the WordInterval a balanced, sorted tree.*) subsection\Syntax\ context notes [[typedef_overloaded]] begin - datatype ('a::len0) wordinterval = WordInterval - "('a::len0) word" \ \start (inclusive)\ - "('a::len0) word" \ \end (inclusive)\ + datatype ('a::len) wordinterval = WordInterval + "('a::len) word" \ \start (inclusive)\ + "('a::len) word" \ \end (inclusive)\ | RangeUnion "'a wordinterval" "'a wordinterval" end subsection\Semantics\ - fun wordinterval_to_set :: "'a::len0 wordinterval \ ('a::len0 word) set" + fun wordinterval_to_set :: "'a::len wordinterval \ ('a::len word) set" where "wordinterval_to_set (WordInterval start end) = {start .. end}" | "wordinterval_to_set (RangeUnion r1 r2) = wordinterval_to_set r1 \ wordinterval_to_set r2" (*Note: The runtime of all the operations could be improved, for example by keeping the tree sorted and balanced.*) subsection\Basic operations\ text\\\\\ - fun wordinterval_element :: "'a::len0 word \ 'a::len0 wordinterval \ bool" where + fun wordinterval_element :: "'a::len word \ 'a::len wordinterval \ bool" where "wordinterval_element el (WordInterval s e) \ s \ el \ el \ e" | "wordinterval_element el (RangeUnion r1 r2) \ wordinterval_element el r1 \ wordinterval_element el r2" lemma wordinterval_element_set_eq[simp]: "wordinterval_element el rg = (el \ wordinterval_to_set rg)" by(induction rg rule: wordinterval_element.induct) simp_all definition wordinterval_union - :: "'a::len0 wordinterval \ 'a::len0 wordinterval \ 'a::len0 wordinterval" where + :: "'a::len wordinterval \ 'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_union r1 r2 = RangeUnion r1 r2" lemma wordinterval_union_set_eq[simp]: "wordinterval_to_set (wordinterval_union r1 r2) = wordinterval_to_set r1 \ wordinterval_to_set r2" unfolding wordinterval_union_def by simp - fun wordinterval_empty :: "'a::len0 wordinterval \ bool" where + fun wordinterval_empty :: "'a::len wordinterval \ bool" where "wordinterval_empty (WordInterval s e) \ e < s" | "wordinterval_empty (RangeUnion r1 r2) \ wordinterval_empty r1 \ wordinterval_empty r2" lemma wordinterval_empty_set_eq[simp]: "wordinterval_empty r \ wordinterval_to_set r = {}" by(induction r) auto definition Empty_WordInterval :: "'a::len wordinterval" where "Empty_WordInterval \ WordInterval 1 0" lemma wordinterval_empty_Empty_WordInterval: "wordinterval_empty Empty_WordInterval" by(simp add: Empty_WordInterval_def) lemma Empty_WordInterval_set_eq[simp]: "wordinterval_to_set Empty_WordInterval = {}" by(simp add: Empty_WordInterval_def) subsection\WordInterval and Lists\ text\A list of \(start, end)\ tuples.\ text\wordinterval to list\ - fun wi2l :: "'a::len0 wordinterval \ ('a::len0 word \ 'a::len0 word) list" where + fun wi2l :: "'a::len wordinterval \ ('a::len word \ 'a::len word) list" where "wi2l (RangeUnion r1 r2) = wi2l r1 @ wi2l r2" | "wi2l (WordInterval s e) = (if e < s then [] else [(s,e)])" text\list to wordinterval\ fun l2wi :: "('a::len word \ 'a word) list \ 'a wordinterval" where "l2wi [] = Empty_WordInterval" | "l2wi [(s,e)] = (WordInterval s e)" | "l2wi ((s,e)#rs) = (RangeUnion (WordInterval s e) (l2wi rs))" lemma l2wi_append: "wordinterval_to_set (l2wi (l1@l2)) = wordinterval_to_set (l2wi l1) \ wordinterval_to_set (l2wi l2)" proof(induction l1 arbitrary: l2 rule:l2wi.induct) case 1 thus ?case by simp next case (2 s e l2) thus ?case by (cases l2) simp_all next case 3 thus ?case by force qed lemma l2wi_wi2l[simp]: "wordinterval_to_set (l2wi (wi2l r)) = wordinterval_to_set r" by(induction r) (simp_all add: l2wi_append) lemma l2wi: "wordinterval_to_set (l2wi l) = (\ (i,j) \ set l. {i .. j})" by(induction l rule: l2wi.induct, simp_all) lemma wi2l: "(\(i,j)\set (wi2l r). {i .. j}) = wordinterval_to_set r" by(induction r rule: wi2l.induct, simp_all) lemma l2wi_remdups[simp]: "wordinterval_to_set (l2wi (remdups ls)) = wordinterval_to_set (l2wi ls)" by(simp add: l2wi) lemma wi2l_empty[simp]: "wi2l Empty_WordInterval = []" unfolding Empty_WordInterval_def by simp subsection\Optimizing and minimizing @{typ "('a::len) wordinterval"}s\ text\Removing empty intervals\ context begin - fun wordinterval_optimize_empty :: "'a::len0 wordinterval \ 'a wordinterval" where + fun wordinterval_optimize_empty :: "'a::len wordinterval \ 'a wordinterval" where "wordinterval_optimize_empty (RangeUnion r1 r2) = (let r1o = wordinterval_optimize_empty r1; r2o = wordinterval_optimize_empty r2 in if wordinterval_empty r1o then r2o else if wordinterval_empty r2o then r1o else RangeUnion r1o r2o)" | "wordinterval_optimize_empty r = r" lemma wordinterval_optimize_empty_set_eq[simp]: "wordinterval_to_set (wordinterval_optimize_empty r) = wordinterval_to_set r" by(induction r) (simp_all add: Let_def) lemma wordinterval_optimize_empty_double: "wordinterval_optimize_empty (wordinterval_optimize_empty r) = wordinterval_optimize_empty r" by(induction r) (simp_all add: Let_def) - private fun wordinterval_empty_shallow :: "'a::len0 wordinterval \ bool" where + private fun wordinterval_empty_shallow :: "'a::len wordinterval \ bool" where "wordinterval_empty_shallow (WordInterval s e) \ e < s" | "wordinterval_empty_shallow (RangeUnion _ _) \ False" private lemma helper_optimize_shallow: "wordinterval_empty_shallow (wordinterval_optimize_empty r) = wordinterval_empty (wordinterval_optimize_empty r)" by(induction r) fastforce+ private fun wordinterval_optimize_empty2 where "wordinterval_optimize_empty2 (RangeUnion r1 r2) = (let r1o = wordinterval_optimize_empty r1; r2o = wordinterval_optimize_empty r2 in if wordinterval_empty_shallow r1o then r2o else if wordinterval_empty_shallow r2o then r1o else RangeUnion r1o r2o)" | "wordinterval_optimize_empty2 r = r" lemma wordinterval_optimize_empty_code[code_unfold]: "wordinterval_optimize_empty = wordinterval_optimize_empty2" by (subst fun_eq_iff, clarify, rename_tac r, induct_tac r) (unfold wordinterval_optimize_empty.simps wordinterval_optimize_empty2.simps Let_def helper_optimize_shallow, simp_all) end text\Merging overlapping intervals\ context begin private definition disjoint :: "'a set \ 'a set \ bool" where "disjoint A B \ A \ B = {}" - private primrec interval_of :: "('a::len0) word \ 'a word \ 'a word set" where + private primrec interval_of :: "('a::len) word \ 'a word \ 'a word set" where "interval_of (s,e) = {s .. e}" declare interval_of.simps[simp del] private definition disjoint_intervals - :: "(('a::len0) word \ ('a::len0) word) \ ('a word \ 'a word) \ bool" + :: "(('a::len) word \ ('a::len) word) \ ('a word \ 'a word) \ bool" where "disjoint_intervals A B \ disjoint (interval_of A) (interval_of B)" private definition not_disjoint_intervals - :: "(('a::len0) word \ ('a::len0) word) \ ('a word \ 'a word) \ bool" + :: "(('a::len) word \ ('a::len) word) \ ('a word \ 'a word) \ bool" where "not_disjoint_intervals A B \ \ disjoint (interval_of A) (interval_of B)" private lemma [code]: "not_disjoint_intervals A B = (case A of (s,e) \ case B of (s',e') \ s \ e' \ s' \ e \ s \ e \ s' \ e')" apply(cases A, cases B) apply(simp add: not_disjoint_intervals_def interval_of.simps disjoint_def) done private lemma [code]: "disjoint_intervals A B = (case A of (s,e) \ case B of (s',e') \ s > e' \ s' > e \ s > e \ s' > e')" apply(cases A, cases B) apply(simp add: disjoint_intervals_def interval_of.simps disjoint_def) by fastforce text\BEGIN merging overlapping intervals\ (*result has no empty intervals and all are disjoint. merging things such as [1,7] [8,10] would still be possible*) private fun merge_overlap - :: "(('a::len0) word \ ('a::len0) word) \ ('a word \ 'a word) list \ ('a word \ 'a word) list" + :: "(('a::len) word \ ('a::len) word) \ ('a word \ 'a word) list \ ('a word \ 'a word) list" where "merge_overlap s [] = [s]" | "merge_overlap (s,e) ((s',e')#ss) = ( if not_disjoint_intervals (s,e) (s',e') then (min s s', max e e')#ss else (s',e')#merge_overlap (s,e) ss)" private lemma not_disjoint_union: - fixes s :: "('a::len0) word" + fixes s :: "('a::len) word" shows "\ disjoint {s..e} {s'..e'} \ {s..e} \ {s'..e'} = {min s s' .. max e e'}" by(auto simp add: disjoint_def min_def max_def) private lemma disjoint_subset: "disjoint A B \ A \ B \ C \ A \ C" unfolding disjoint_def by blast private lemma merge_overlap_helper1: "interval_of A \ (\s \ set ss. interval_of s) \ (\s \ set (merge_overlap A ss). interval_of s) = (\s \ set ss. interval_of s)" apply(induction ss) apply(simp; fail) apply(rename_tac x xs) apply(cases A, rename_tac a b) apply(case_tac x) apply(simp add: not_disjoint_intervals_def interval_of.simps) apply(intro impI conjI) apply(drule not_disjoint_union) apply blast apply(drule_tac C="(\x\set xs. interval_of x)" in disjoint_subset) apply(simp_all) done private lemma merge_overlap_helper2: "\s'\set ss. \ disjoint (interval_of A) (interval_of s') \ interval_of A \ (\s \ set ss. interval_of s) = (\s \ set (merge_overlap A ss). interval_of s)" apply(induction ss) apply(simp; fail) apply(rename_tac x xs) apply(cases A, rename_tac a b) apply(case_tac x) apply(simp add: not_disjoint_intervals_def interval_of.simps) apply(intro impI conjI) apply(drule not_disjoint_union) apply blast apply(simp) by blast private lemma merge_overlap_length: "\s' \ set ss. \ disjoint (interval_of A) (interval_of s') \ length (merge_overlap A ss) = length ss" apply(induction ss) apply(simp) apply(rename_tac x xs) apply(cases A, rename_tac a b) apply(case_tac x) apply(simp add: not_disjoint_intervals_def interval_of.simps) done lemma "merge_overlap (1:: 16 word,2) [(1, 7)] = [(1, 7)]" by eval lemma "merge_overlap (1:: 16 word,2) [(2, 7)] = [(1, 7)]" by eval lemma "merge_overlap (1:: 16 word,2) [(3, 7)] = [(3, 7), (1,2)]" by eval private function listwordinterval_compress - :: "(('a::len0) word \ ('a::len0) word) list \ ('a word \ 'a word) list" where + :: "(('a::len) word \ ('a::len) word) list \ ('a word \ 'a word) list" where "listwordinterval_compress [] = []" | "listwordinterval_compress (s#ss) = ( if \s' \ set ss. disjoint_intervals s s' then s#listwordinterval_compress ss else listwordinterval_compress (merge_overlap s ss))" by(pat_completeness, auto) private termination listwordinterval_compress apply (relation "measure length") apply(rule wf_measure) apply(simp) using disjoint_intervals_def merge_overlap_length by fastforce private lemma listwordinterval_compress: "(\s \ set (listwordinterval_compress ss). interval_of s) = (\s \ set ss. interval_of s)" apply(induction ss rule: listwordinterval_compress.induct) apply(simp) apply(simp) apply(intro impI) apply(simp add: disjoint_intervals_def) apply(drule merge_overlap_helper2) apply(simp) done lemma "listwordinterval_compress [(1::32 word,3), (8,10), (2,5), (3,7)] = [(8, 10), (1, 7)]" by eval private lemma A_in_listwordinterval_compress: "A \ set (listwordinterval_compress ss) \ interval_of A \ (\s \ set ss. interval_of s)" using listwordinterval_compress by blast private lemma listwordinterval_compress_disjoint: "A \ set (listwordinterval_compress ss) \ B \ set (listwordinterval_compress ss) \ A \ B \ disjoint (interval_of A) (interval_of B)" apply(induction ss arbitrary: rule: listwordinterval_compress.induct) apply(simp) apply(simp split: if_split_asm) apply(elim disjE) apply(simp_all) apply(simp_all add: disjoint_intervals_def disjoint_def) apply(blast dest: A_in_listwordinterval_compress)+ done text\END merging overlapping intervals\ text\BEGIN merging adjacent intervals\ private fun merge_adjacent :: "(('a::len) word \ ('a::len) word) \ ('a word \ 'a word) list \ ('a word \ 'a word) list" where "merge_adjacent s [] = [s]" | "merge_adjacent (s,e) ((s',e')#ss) = ( if s \e \ s' \ e' \ word_next e = s' then (s, e')#ss else if s \e \ s' \ e' \ word_next e' = s then (s', e)#ss else (s',e')#merge_adjacent (s,e) ss)" private lemma merge_adjacent_helper: "interval_of A \ (\s \ set ss. interval_of s) = (\s \ set (merge_adjacent A ss). interval_of s)" apply(induction ss) apply(simp; fail) apply(rename_tac x xs) apply(cases A, rename_tac a b) apply(case_tac x) apply(simp add: interval_of.simps) apply(intro impI conjI) apply (metis Un_assoc word_adjacent_union) apply(elim conjE) apply(drule(2) word_adjacent_union) subgoal by (blast) subgoal by (metis word_adjacent_union Un_assoc) by blast private lemma merge_adjacent_length: "\(s', e')\set ss. s \ e \ s' \ e' \ (word_next e = s' \ word_next e' = s) \ length (merge_adjacent (s,e) ss) = length ss" apply(induction ss) apply(simp) apply(rename_tac x xs) apply(case_tac x) apply(simp add: ) by blast private function listwordinterval_adjacent :: "(('a::len) word \ ('a::len) word) list \ ('a word \ 'a word) list" where "listwordinterval_adjacent [] = []" | "listwordinterval_adjacent ((s,e)#ss) = ( if \(s',e') \ set ss. \ (s \e \ s' \ e' \ (word_next e = s' \ word_next e' = s)) then (s,e)#listwordinterval_adjacent ss else listwordinterval_adjacent (merge_adjacent (s,e) ss))" by(pat_completeness, auto) private termination listwordinterval_adjacent apply (relation "measure length") apply(rule wf_measure) apply(simp) apply(simp) using merge_adjacent_length by fastforce private lemma listwordinterval_adjacent: "(\s \ set (listwordinterval_adjacent ss). interval_of s) = (\s \ set ss. interval_of s)" apply(induction ss rule: listwordinterval_adjacent.induct) apply(simp) apply(simp add: merge_adjacent_helper) done lemma "listwordinterval_adjacent [(1::16 word, 3), (5, 10), (10,10), (4,4)] = [(10, 10), (1, 10)]" by eval text\END merging adjacent intervals\ definition wordinterval_compress :: "('a::len) wordinterval \ 'a wordinterval" where "wordinterval_compress r \ l2wi (remdups (listwordinterval_adjacent (listwordinterval_compress (wi2l (wordinterval_optimize_empty r)))))" text\Correctness: Compression preserves semantics\ lemma wordinterval_compress: "wordinterval_to_set (wordinterval_compress r) = wordinterval_to_set r" unfolding wordinterval_compress_def proof - have interval_of': "interval_of s = (case s of (s,e) \ {s .. e})" for s by (cases s) (simp add: interval_of.simps) have "wordinterval_to_set (l2wi (remdups (listwordinterval_adjacent (listwordinterval_compress (wi2l (wordinterval_optimize_empty r)))))) = (\x\set (listwordinterval_adjacent (listwordinterval_compress (wi2l (wordinterval_optimize_empty r)))). interval_of x)" by (force simp: interval_of' l2wi) also have "\ = (\s\set (wi2l (wordinterval_optimize_empty r)). interval_of s)" by(simp add: listwordinterval_compress listwordinterval_adjacent) also have "\ = (\(i, j)\set (wi2l (wordinterval_optimize_empty r)). {i..j})" by(simp add: interval_of') also have "\ = wordinterval_to_set r" by(simp add: wi2l) finally show "wordinterval_to_set (l2wi (remdups (listwordinterval_adjacent (listwordinterval_compress (wi2l (wordinterval_optimize_empty r)))))) = wordinterval_to_set r" . qed end text\Example\ lemma "(wi2l \ (wordinterval_compress :: 32 wordinterval \ 32 wordinterval) \ l2wi) [(70, 80001), (0,0), (150, 8000), (1,3), (42,41), (3,7), (56, 200), (8,10)] = [(56, 80001), (0, 10)]" by eval lemma "wordinterval_compress (RangeUnion (RangeUnion (WordInterval (1::32 word) 5) (WordInterval 8 10)) (WordInterval 3 7)) = WordInterval 1 10" by eval subsection\Further operations\ text\\\\\ definition wordinterval_Union :: "('a::len) wordinterval list \ 'a wordinterval" where "wordinterval_Union ws = wordinterval_compress (foldr wordinterval_union ws Empty_WordInterval)" lemma wordinterval_Union: "wordinterval_to_set (wordinterval_Union ws) = (\ w \ (set ws). wordinterval_to_set w)" by(induction ws) (simp_all add: wordinterval_compress wordinterval_Union_def) context begin private fun wordinterval_setminus' :: "'a::len wordinterval \ 'a wordinterval \ 'a wordinterval" where "wordinterval_setminus' (WordInterval s e) (WordInterval ms me) = ( if s > e \ ms > me then WordInterval s e else if me \ e then WordInterval (if ms = 0 then 1 else s) (min e (word_prev ms)) else if ms \ s then WordInterval (max s (word_next me)) (if me = max_word then 0 else e) else RangeUnion (WordInterval (if ms = 0 then 1 else s) (word_prev ms)) (WordInterval (word_next me) (if me = max_word then 0 else e)) )" | "wordinterval_setminus' (RangeUnion r1 r2) t = RangeUnion (wordinterval_setminus' r1 t) (wordinterval_setminus' r2 t)"| "wordinterval_setminus' t (RangeUnion r1 r2) = wordinterval_setminus' (wordinterval_setminus' t r1) r2" private lemma wordinterval_setminus'_rr_set_eq: "wordinterval_to_set(wordinterval_setminus' (WordInterval s e) (WordInterval ms me)) = wordinterval_to_set (WordInterval s e) - wordinterval_to_set (WordInterval ms me)" apply(simp only: wordinterval_setminus'.simps) apply(case_tac "e < s") apply simp apply(case_tac "me < ms") apply simp apply(case_tac [!] "e \ me") apply(case_tac [!] "ms = 0") apply(case_tac [!] "ms \ s") apply(case_tac [!] "me = max_word") apply(simp_all add: word_prev_def word_next_def min_def max_def) apply(safe) apply(auto) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) apply(uint_arith) done private lemma wordinterval_setminus'_set_eq: "wordinterval_to_set (wordinterval_setminus' r1 r2) = wordinterval_to_set r1 - wordinterval_to_set r2" apply(induction rule: wordinterval_setminus'.induct) using wordinterval_setminus'_rr_set_eq apply blast apply auto done lemma wordinterval_setminus'_empty_struct: "wordinterval_empty r2 \ wordinterval_setminus' r1 r2 = r1" by(induction r1 r2 rule: wordinterval_setminus'.induct) auto definition wordinterval_setminus :: "'a::len wordinterval \ 'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_setminus r1 r2 = wordinterval_compress (wordinterval_setminus' r1 r2)" lemma wordinterval_setminus_set_eq[simp]: "wordinterval_to_set (wordinterval_setminus r1 r2) = wordinterval_to_set r1 - wordinterval_to_set r2" by(simp add: wordinterval_setminus_def wordinterval_compress wordinterval_setminus'_set_eq) end definition wordinterval_UNIV :: "'a::len wordinterval" where "wordinterval_UNIV \ WordInterval 0 max_word" lemma wordinterval_UNIV_set_eq[simp]: "wordinterval_to_set wordinterval_UNIV = UNIV" unfolding wordinterval_UNIV_def using max_word_max by fastforce fun wordinterval_invert :: "'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_invert r = wordinterval_setminus wordinterval_UNIV r" lemma wordinterval_invert_set_eq[simp]: "wordinterval_to_set (wordinterval_invert r) = UNIV - wordinterval_to_set r" by(auto) lemma wordinterval_invert_UNIV_empty: "wordinterval_empty (wordinterval_invert wordinterval_UNIV)" by simp lemma wi2l_univ[simp]: "wi2l wordinterval_UNIV = [(0, max_word)]" unfolding wordinterval_UNIV_def by simp text\\\\\ context begin private lemma "{(s::nat) .. e} \ {s' .. e'} = {} \ s > e' \ s' > e \ s > e \ s' > e'" by simp linarith private fun wordinterval_intersection' :: "'a::len wordinterval \ 'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_intersection' (WordInterval s e) (WordInterval s' e') = ( if s > e \ s' > e' \ s > e' \ s' > e \ s > e \ s' > e' then Empty_WordInterval else WordInterval (max s s') (min e e') )" | "wordinterval_intersection' (RangeUnion r1 r2) t = RangeUnion (wordinterval_intersection' r1 t) (wordinterval_intersection' r2 t)"| "wordinterval_intersection' t (RangeUnion r1 r2) = RangeUnion (wordinterval_intersection' t r1) (wordinterval_intersection' t r2)" private lemma wordinterval_intersection'_set_eq: "wordinterval_to_set (wordinterval_intersection' r1 r2) = wordinterval_to_set r1 \ wordinterval_to_set r2" by(induction r1 r2 rule: wordinterval_intersection'.induct) (auto) lemma "wordinterval_intersection' (RangeUnion (RangeUnion (WordInterval (1::32 word) 3) (WordInterval 8 10)) (WordInterval 1 3)) (WordInterval 1 3) = RangeUnion (RangeUnion (WordInterval 1 3) (WordInterval 1 0)) (WordInterval 1 3)" by eval definition wordinterval_intersection :: "'a::len wordinterval \ 'a::len wordinterval \ 'a::len wordinterval" where "wordinterval_intersection r1 r2 \ wordinterval_compress (wordinterval_intersection' r1 r2)" lemma wordinterval_intersection_set_eq[simp]: "wordinterval_to_set (wordinterval_intersection r1 r2) = wordinterval_to_set r1 \ wordinterval_to_set r2" by(simp add: wordinterval_intersection_def wordinterval_compress wordinterval_intersection'_set_eq) lemma "wordinterval_intersection (RangeUnion (RangeUnion (WordInterval (1::32 word) 3) (WordInterval 8 10)) (WordInterval 1 3)) (WordInterval 1 3) = WordInterval 1 3" by eval end definition wordinterval_subset :: "'a::len wordinterval \ 'a::len wordinterval \ bool" where "wordinterval_subset r1 r2 \ wordinterval_empty (wordinterval_setminus r1 r2)" lemma wordinterval_subset_set_eq[simp]: "wordinterval_subset r1 r2 = (wordinterval_to_set r1 \ wordinterval_to_set r2)" unfolding wordinterval_subset_def by simp definition wordinterval_eq :: "'a::len wordinterval \ 'a::len wordinterval \ bool" where "wordinterval_eq r1 r2 = (wordinterval_subset r1 r2 \ wordinterval_subset r2 r1)" lemma wordinterval_eq_set_eq: "wordinterval_eq r1 r2 \ wordinterval_to_set r1 = wordinterval_to_set r2" unfolding wordinterval_eq_def by auto thm iffD1[OF wordinterval_eq_set_eq] (*declare iffD1[OF wordinterval_eq_set_eq, simp]*) lemma wordinterval_eq_comm: "wordinterval_eq r1 r2 \ wordinterval_eq r2 r1" unfolding wordinterval_eq_def by fast lemma wordinterval_to_set_alt: "wordinterval_to_set r = {x. wordinterval_element x r}" unfolding wordinterval_element_set_eq by blast lemma wordinterval_un_empty: "wordinterval_empty r1 \ wordinterval_eq (wordinterval_union r1 r2) r2" by(subst wordinterval_eq_set_eq, simp) lemma wordinterval_un_emty_b: "wordinterval_empty r2 \ wordinterval_eq (wordinterval_union r1 r2) r1" by(subst wordinterval_eq_set_eq, simp) lemma wordinterval_Diff_triv: "wordinterval_empty (wordinterval_intersection a b) \ wordinterval_eq (wordinterval_setminus a b) a" unfolding wordinterval_eq_set_eq by simp blast text\A size of the datatype, does not correspond to the cardinality of the corresponding set\ fun wordinterval_size :: "('a::len) wordinterval \ nat" where "wordinterval_size (RangeUnion a b) = wordinterval_size a + wordinterval_size b" | "wordinterval_size (WordInterval s e) = (if s \ e then 1 else 0)" lemma wordinterval_size_length: "wordinterval_size r = length (wi2l r)" by(induction r) (auto) lemma Ex_wordinterval_nonempty: "\x::('a::len wordinterval). y \ wordinterval_to_set x" proof show "y \ wordinterval_to_set wordinterval_UNIV" by simp qed lemma wordinterval_eq_reflp: "reflp wordinterval_eq" apply(rule reflpI) by(simp only: wordinterval_eq_set_eq) lemma wordintervalt_eq_symp: "symp wordinterval_eq" apply(rule sympI) by(simp add: wordinterval_eq_comm) lemma wordinterval_eq_transp: "transp wordinterval_eq" apply(rule transpI) by(simp only: wordinterval_eq_set_eq) lemma wordinterval_eq_equivp: "equivp wordinterval_eq" by (auto intro: equivpI wordinterval_eq_reflp wordintervalt_eq_symp wordinterval_eq_transp) text\The smallest element in the interval\ definition is_lowest_element :: "'a::ord \ 'a set \ bool" where "is_lowest_element x S = (x \ S \ (\y\S. y \ x \ y = x))" lemma fixes x :: "'a :: complete_lattice" assumes "x \ S" shows " x = Inf S \ is_lowest_element x S" using assms apply(simp add: is_lowest_element_def) by (simp add: Inf_lower eq_iff) lemma fixes x :: "'a :: linorder" assumes "finite S" and "x \ S" shows "is_lowest_element x S \ x = Min S" apply(rule) subgoal apply(simp add: is_lowest_element_def) apply(subst Min_eqI[symmetric]) using assms by(auto) by (metis Min.coboundedI assms(1) assms(2) dual_order.antisym is_lowest_element_def) text\Smallest element in the interval\ - fun wordinterval_lowest_element :: "'a::len0 wordinterval \ 'a word option" where + fun wordinterval_lowest_element :: "'a::len wordinterval \ 'a word option" where "wordinterval_lowest_element (WordInterval s e) = (if s \ e then Some s else None)" | "wordinterval_lowest_element (RangeUnion A B) = (case (wordinterval_lowest_element A, wordinterval_lowest_element B) of (Some a, Some b) \ Some (if a < b then a else b) | (None, Some b) \ Some b | (Some a, None) \ Some a | (None, None) \ None)" lemma wordinterval_lowest_none_empty: "wordinterval_lowest_element r = None \ wordinterval_empty r" proof(induction r) case WordInterval thus ?case by simp next case RangeUnion thus ?case by fastforce qed lemma wordinterval_lowest_element_correct_A: "wordinterval_lowest_element r = Some x \ is_lowest_element x (wordinterval_to_set r)" unfolding is_lowest_element_def apply(induction r arbitrary: x rule: wordinterval_lowest_element.induct) apply(rename_tac rs re x, case_tac "rs \ re", auto)[1] apply(subst(asm) wordinterval_lowest_element.simps(2)) apply(rename_tac A B x) apply(case_tac "wordinterval_lowest_element B") apply(case_tac[!] "wordinterval_lowest_element A") apply(simp_all add: wordinterval_lowest_none_empty)[3] apply fastforce done lemma wordinterval_lowest_element_set_eq: assumes "\ wordinterval_empty r" shows "(wordinterval_lowest_element r = Some x) = (is_lowest_element x (wordinterval_to_set r))" (*unfolding is_lowest_element_def*) proof(rule iffI) assume "wordinterval_lowest_element r = Some x" thus "is_lowest_element x (wordinterval_to_set r)" using wordinterval_lowest_element_correct_A wordinterval_lowest_none_empty by simp next assume "is_lowest_element x (wordinterval_to_set r)" with assms show "(wordinterval_lowest_element r = Some x)" proof(induction r arbitrary: x rule: wordinterval_lowest_element.induct) case 1 thus ?case by(simp add: is_lowest_element_def) next case (2 A B x) have is_lowest_RangeUnion: "is_lowest_element x (wordinterval_to_set A \ wordinterval_to_set B) \ is_lowest_element x (wordinterval_to_set A) \ is_lowest_element x (wordinterval_to_set B)" by(simp add: is_lowest_element_def) (*why \ A B?*) have wordinterval_lowest_element_RangeUnion: "\a b A B. wordinterval_lowest_element A = Some a \ wordinterval_lowest_element B = Some b \ wordinterval_lowest_element (RangeUnion A B) = Some (min a b)" by(auto dest!: wordinterval_lowest_element_correct_A simp add: is_lowest_element_def min_def) from 2 show ?case apply(case_tac "wordinterval_lowest_element B") apply(case_tac[!] "wordinterval_lowest_element A") apply(auto simp add: is_lowest_element_def)[3] apply(subgoal_tac "\ wordinterval_empty A \ \ wordinterval_empty B") prefer 2 using arg_cong[where f = Not, OF wordinterval_lowest_none_empty] apply blast apply(drule(1) wordinterval_lowest_element_RangeUnion) apply(simp split: option.split_asm add: min_def) apply(drule is_lowest_RangeUnion) apply(elim disjE) apply(simp add: is_lowest_element_def) apply(clarsimp simp add: wordinterval_lowest_none_empty) apply(simp add: is_lowest_element_def) apply(clarsimp simp add: wordinterval_lowest_none_empty) using wordinterval_lowest_element_correct_A[simplified is_lowest_element_def] by (metis Un_iff not_le) qed qed text\Cardinality approximation for @{typ "('a::len) wordinterval"}s\ context begin lemma card_atLeastAtMost_word: fixes s::"('a::len) word" shows "card {s..e} = Suc (unat e) - (unat s)" apply(cases "s > e") apply(simp) apply(subst(asm) Word.word_less_nat_alt) apply simp apply(subst Word_Lemmas.upto_enum_set_conv2[symmetric]) apply(subst List.card_set) apply(simp add: remdups_enum_upto) done fun wordinterval_card :: "('a::len) wordinterval \ nat" where "wordinterval_card (WordInterval s e) = Suc (unat e) - (unat s)" | "wordinterval_card (RangeUnion a b) = wordinterval_card a + wordinterval_card b" lemma wordinterval_card: "wordinterval_card r \ card (wordinterval_to_set r)" proof(induction r) case WordInterval thus ?case by (simp add: card_atLeastAtMost_word) next case (RangeUnion r1 r2) have "card (wordinterval_to_set r1 \ wordinterval_to_set r2) \ card (wordinterval_to_set r1) + card (wordinterval_to_set r2)" using Finite_Set.card_Un_le by blast with RangeUnion show ?case by(simp) qed text\With @{thm wordinterval_compress} it should be possible to get the exact cardinality\ end end diff --git a/thys/Interval_Arithmetic_Word32/Interval_Word32.thy b/thys/Interval_Arithmetic_Word32/Interval_Word32.thy --- a/thys/Interval_Arithmetic_Word32/Interval_Word32.thy +++ b/thys/Interval_Arithmetic_Word32/Interval_Word32.thy @@ -1,3880 +1,3885 @@ (* Author: Brandon Bohrer *) theory Interval_Word32 imports Complex_Main Word_Lib.Word_Lemmas Word_Lib.Word_Lib Word_Lib.Word_Syntax begin text\Interval-Word32.thy implements conservative interval arithmetic operators on 32-bit word values, with explicit infinities for values outside the representable bounds. It is suitable for use in interpreters for languages which must have a well-understood low-level behavior (see Interpreter.thy). This work was originally part of the paper by Bohrer \emph{et al.}~\cite{BohrerTMMP18}. It is worth noting that this is not the first formalization of interval arithmetic in Isabelle/HOL. This article is presented regardless because it has unique goals in mind which have led to unique design decisions. Our goal is generate code which can be used to perform conservative arithmetic in implementations extracted from a proof. The Isabelle standard library now features interval arithmetic, for example in Approximation.thy. Ours differs in two ways: 1) We use intervals with explicit positive and negative infinities, and with overflow checking. Such checking is often relevant in implementation-level code with unknown inputs. To promote memory-efficient implementations, we moreover use sentinel values for infinities, rather than datatype constructors. This is especially important in real-time settings where the garbarge collection required for datatypes can be a concern. 2) Our goal is not to use interval arithmetic to discharge Isabelle goals, but to generate useful proven-correct implementation code, see Interpreter.thy. On the other hand, we are not concerned with producing interval-based automation for arithmetic goals in HOL. In practice, much of the work in this theory comes down to sheer case-analysis. Bounds-checking requires many edge cases in arithmetic functions, which come with many cases in proofs. Where possible, we attempt to offload interesting facts about word representations of numbers into reusable lemmas, but even then main results require many subcases, each with a certain amount of arithmetic grunt work. \ section \Interval arithmetic definitions\ subsection \Syntax\ text\Words are 32-bit\ type_synonym word = "32 Word.word" text\Sentinel values for infinities. Note that we leave the maximum value ($2^31$) completed unused, so that negation of $(2^{31})-1$ is not an edge case\ definition NEG_INF::"word" where NEG_INF_def[simp]:"NEG_INF = -((2 ^ 31) -1)" definition NegInf::"real" where NegInf[simp]:"NegInf = real_of_int (sint NEG_INF)" definition POS_INF::"word" where POS_INF_def[simp]:"POS_INF = (2^31) - 1" definition PosInf::"real" where PosInf[simp]:"PosInf = real_of_int (sint POS_INF)" text\Subtype of words who represent a finite value. \ typedef bword = "{n::word. sint n \ sint NEG_INF \ sint n \ sint POS_INF}" apply(rule exI[where x=NEG_INF]) by (auto) text\Numeric literals\ type_synonym lit = bword setup_lifting type_definition_bword lift_definition bword_zero::"bword" is "0::32 Word.word" by auto lift_definition bword_one::"bword" is "1::32 Word.word" by(auto simp add: sint_uint) lift_definition bword_neg_one::"bword" is "-1::32 Word.word" by(auto) definition word::"word \ bool" where word_def[simp]:"word w \ w \ {NEG_INF..POS_INF}" named_theorems rep_simps "Simplifications for representation functions" text\Definitions of interval containment and word representation repe w r iff word w encodes real number r\ inductive repe ::"word \ real \ bool" (infix "\\<^sub>E" 10) where repPOS_INF:"r \ real_of_int (sint POS_INF) \ repe POS_INF r" | repNEG_INF:"r \ real_of_int (sint NEG_INF) \ repe NEG_INF r" | repINT:"(sint w) < real_of_int(sint POS_INF) \ (sint w) > real_of_int(sint NEG_INF) \ repe w (sint w)" inductive_simps repePos_simps[rep_simps]:"repe POS_INF r" and repeNeg_simps[rep_simps]:"repe NEG_INF r" and repeInt_simps[rep_simps]:"repe w (sint w)" text\repU w r if w represents an upper bound of r\ definition repU ::"word \ real \ bool" (infix "\\<^sub>U" 10) where "repU w r \ \ r'. r' \ r \ repe w r'" lemma repU_leq:"repU w r \ r' \ r \ repU w r'" unfolding repU_def using order_trans by auto text\repU w r if w represents a lower bound of r\ definition repL ::"word \ real \ bool" (infix "\\<^sub>L" 10) where "repL w r \ \ r'. r' \ r \ repe w r'" lemma repL_geq:"repL w r \ r' \ r \ repL w r'" unfolding repL_def using order_trans by auto text\repP (l,u) r iff l and u encode lower and upper bounds of r\ definition repP ::"word * word \ real \ bool" (infix "\\<^sub>P" 10) where "repP w r \ let (w1, w2) = w in repL w1 r \ repU w2 r" lemma int_not_posinf: assumes b1:"real_of_int (sint ra) < real_of_int (sint POS_INF)" assumes b2:"real_of_int (sint NEG_INF) < real_of_int (sint ra)" shows "ra \ POS_INF" using b1 b2 by auto lemma int_not_neginf: assumes b1:" real_of_int (sint ra) < real_of_int (sint POS_INF)" assumes b2:" real_of_int (sint NEG_INF) < real_of_int (sint ra)" shows "ra \ NEG_INF" using b1 b2 by auto lemma int_not_undef: assumes b1:"real_of_int (sint ra) < real_of_int (sint POS_INF)" assumes b2:"real_of_int (sint NEG_INF) < real_of_int (sint ra)" shows "ra \ NEG_INF-1" using b1 b2 by auto lemma sint_range: assumes b1:"real_of_int (sint ra) < real_of_int (sint POS_INF)" assumes b2:"real_of_int (sint NEG_INF) < real_of_int (sint ra)" shows "sint ra \ {i. i > -((2^31)-1) \ i < (2^31)-1}" using b1 b2 by auto lemma word_size_neg: fixes w :: "32 Word.word" shows "size (-w) = size w" using Word.word_size[of w] Word.word_size[of "-w"] by auto lemma uint_distinct: fixes w1 w2 shows "w1 \ w2 \ uint w1 \ uint w2" by auto section \Preliminary lemmas\ subsection \Case analysis lemmas\ text\Case analysis principle for pairs of intervals, used in proofs of arithmetic operations\ lemma ivl_zero_case: fixes l1 u1 l2 u2 :: real assumes ivl1:"l1 \ u1" assumes ivl2:"l2 \ u2" shows "(l1 \ 0 \ 0 \ u1 \ l2 \ 0 \ 0 \ u2) \(l1 \ 0 \ 0 \ u1 \ 0 \ l2) \(l1 \ 0 \ 0 \ u1 \ u2 \ 0) \(0 \ l1 \ l2 \ 0 \ 0 \ u2) \(u1 \ 0 \ l2 \ 0 \ 0 \ u2) \(u1 \ 0 \ u2 \ 0) \(u1 \ 0 \ 0 \ l2) \(0 \ l1 \ u2 \ 0) \(0 \ l1 \ 0 \ l2)" using ivl1 ivl2 by (metis le_cases) lemma case_ivl_zero [consumes 2, case_names ZeroZero ZeroPos ZeroNeg PosZero NegZero NegNeg NegPos PosNeg PosPos]: fixes l1 u1 l2 u2 :: real shows "l1 \ u1 \ l2 \ u2 \ ((l1 \ 0 \ 0 \ u1 \ l2 \ 0 \ 0 \ u2) \ P) \ ((l1 \ 0 \ 0 \ u1 \ 0 \ l2) \ P) \ ((l1 \ 0 \ 0 \ u1 \ u2 \ 0) \ P) \ ((0 \ l1 \ l2 \ 0 \ 0 \ u2) \ P) \ ((u1 \ 0 \ l2 \ 0 \ 0 \ u2) \ P) \ ((u1 \ 0 \ u2 \ 0) \ P) \ ((u1 \ 0 \ 0 \ l2) \ P) \ ((0 \ l1 \ u2 \ 0) \ P) \ ((0 \ l1 \ 0 \ l2) \ P) \ P" using ivl_zero_case[of l1 u1 l2 u2] by auto lemma case_inf2[case_names PosPos PosNeg PosNum NegPos NegNeg NegNum NumPos NumNeg NumNum]: shows "\w1 w2 P. (w1 = POS_INF \ w2 = POS_INF \ P w1 w2) \ (w1 = POS_INF \ w2 = NEG_INF \ P w1 w2) \ (w1 = POS_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ (w1 = NEG_INF \ w2 = POS_INF \ P w1 w2) \ (w1 = NEG_INF \ w2 = NEG_INF \ P w1 w2) \ (w1 = NEG_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 = POS_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 = NEG_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ P w1 w2" by(auto) lemma case_pu_inf[case_names PosAny AnyPos NegNeg NegNum NumNeg NumNum]: shows "\w1 w2 P. (w1 = POS_INF \ P w1 w2) \ (w2 = POS_INF \ P w1 w2) \ (w1 = NEG_INF \ w2 = NEG_INF \ P w1 w2) \ (w1 = NEG_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 = NEG_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ P w1 w2" by(auto) lemma case_pl_inf[case_names NegAny AnyNeg PosPos PosNum NumPos NumNum]: shows "\w1 w2 P. (w1 = NEG_INF \ P w1 w2) \ (w2 = NEG_INF \ P w1 w2) \ (w1 = POS_INF \ w2 = POS_INF \ P w1 w2) \ (w1 = POS_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 = POS_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ P w1 w2" by(auto) lemma word_trichotomy[case_names Less Equal Greater]: fixes w1 w2 :: word shows "(w1 P w1 w2) \ (w1 = w2 \ P w1 w2) \ (w2 P w1 w2) \ P w1 w2" using signed.linorder_cases by auto lemma case_times_inf [case_names PosPos NegPos PosNeg NegNeg PosLo PosHi PosZero NegLo NegHi NegZero LoPos HiPos ZeroPos LoNeg HiNeg ZeroNeg AllFinite]: fixes w1 w2 P assumes pp:"(w1 = POS_INF \ w2 = POS_INF \ P w1 w2)" and np:"(w1 = NEG_INF \ w2 = POS_INF \ P w1 w2)" and pn:"(w1 = POS_INF \ w2 = NEG_INF \ P w1 w2)" and nn:"(w1 = NEG_INF \ w2 = NEG_INF \ P w1 w2)" and pl:"(w1 = POS_INF \ w2 \ NEG_INF \ w2 P w1 w2)" and ph:"(w1 = POS_INF \ w2 \ POS_INF \ 0 P w1 w2)" and pz:"(w1 = POS_INF \ w2 = 0 \ P w1 w2)" and nl:"(w1 = NEG_INF \ w2 \ NEG_INF \ w2 P w1 w2)" and nh:"(w1 = NEG_INF \ w2 \ POS_INF \ 0 P w1 w2)" and nz:"(w1 = NEG_INF \ 0 = w2 \ P w1 w2)" and lp:"(w1 \ NEG_INF \ w1 w2 = POS_INF \ P w1 w2)" and hp:"(w1 \ POS_INF \ 0 w2 = POS_INF \ P w1 w2)" and zp:"(0 = w1 \ w2 = POS_INF \ P w1 w2)" and ln:"(w1 \ NEG_INF \ w1 w2 = NEG_INF \ P w1 w2)" and hn:"(w1 \ POS_INF \ 0 w2 = NEG_INF \ P w1 w2)" and zn:"(0 = w1 \ w2 = NEG_INF \ P w1 w2)" and allFinite:"w1 \ NEG_INF \ w1 \ POS_INF \ w2 \ NEG_INF \ w2 \ POS_INF \ P w1 w2" shows " P w1 w2" proof (cases rule: word_trichotomy[of w1 0]) case Less then have w1l:"w1 Trivial arithmetic lemmas\ lemma max_diff_pos:"0 \ 9223372034707292161 + ((-(2 ^ 31))::real)" by auto lemma max_less:"2 ^ 31 < (9223372039002259455::int)" by auto lemma sints64:"sints 64 = {i. - (2 ^ 63) \ i \ i < 2 ^ 63}" using sints_def[of 64] range_sbintrunc[of 63] by auto lemma sints32:"sints 32 = {i. - (2 ^ 31) \ i \ i < 2 ^ 31}" using sints_def[of 32] range_sbintrunc[of 31] by auto lemma upcast_max:"sint((scast(0x80000001::word))::64 Word.word)=sint((0x80000001::32 Word.word))" by auto lemma upcast_min:"(0xFFFFFFFF80000001::64 Word.word) = ((scast (-0x7FFFFFFF::word))::64 Word.word)" by auto lemma min_extend_neg:"sint ((0xFFFFFFFF80000001)::64 Word.word) < 0" by auto lemma min_extend_val':"sint ((-0x7FFFFFFF)::64 Word.word) = (-0x7FFFFFFF)" by auto lemma min_extend_val:"(-0x7FFFFFFF::64 Word.word) = 0xFFFFFFFF80000001" by auto lemma range2s:"\x::int. x \ 2 ^ 31 - 1 \ x + (- 2147483647) < 2147483647" by auto section \Arithmetic operations\ text\This section defines operations which conservatively compute upper and lower bounds of arithmetic functions given upper and lower bounds on their arguments. Each function comes with a proof that it rounds in the advertised direction. \ subsection \Addition upper bound\ text\Upper bound of w1 + w2\ fun pu :: "word \ word \ word" where "pu w1 w2 = (if w1 = POS_INF then POS_INF else if w2 = POS_INF then POS_INF else if w1 = NEG_INF then (if w2 = NEG_INF then NEG_INF else (let sum::64 Word.word = ((scast w2)::64 Word.word) + ((scast NEG_INF)::64 Word.word) in if (sum::64 Word.word) <=s ((scast NEG_INF)::64 Word.word) then NEG_INF else scast sum)) else if w2 = NEG_INF then (let sum::64 Word.word = ((scast w1)::64 Word.word) + ((scast NEG_INF)::64 Word.word) in if (sum::64 Word.word) <=s ((scast NEG_INF)::64 Word.word) then NEG_INF else scast sum) else (let sum::64 Word.word = ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word) in if ((scast POS_INF)::64 Word.word) <=s (sum::64 Word.word) then POS_INF else if (sum::64 Word.word) <=s ((scast NEG_INF)::64 Word.word) then NEG_INF else scast sum))" lemma scast_down_range: fixes w::"'a::len Word.word" assumes "sint w \ sints (len_of (TYPE('b::len)))" shows "sint w = sint ((scast w)::'b Word.word)" unfolding scast_def by (simp add: assms word_sint.Abs_inverse) lemma pu_lemma: fixes w1 w2 fixes r1 r2 :: real assumes up1:"w1 \\<^sub>U (r1::real)" assumes up2:"w2 \\<^sub>U (r2::real)" shows "pu w1 w2 \\<^sub>U (r1 + r2)" proof - have scast_eq1:"sint((scast w1)::64 Word.word) = sint w1" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have scast_eq2:"sint((scast (0x80000001::word))::64 Word.word) = sint ((0x80000001::32 Word.word))" by auto have scast_eq3:"sint((scast w2)::64 Word.word) = sint w2" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have w2Geq:"sint ((scast w2)::64 Word.word) \ - (2 ^ 31) " using Word.word_sint.Rep[of "(w2)::32 Word.word"] sints32 Word.word_size scast_eq1 upcast_max scast_eq3 len32 mem_Collect_eq by auto have "sint ((scast w2)::64 Word.word) \ 2 ^ 31" apply (auto simp add: Word.word_sint.Rep[of "(w2)::32 Word.word"] sints32 scast_eq3 len32) using Word.word_sint.Rep[of "(w2)::32 Word.word"] len32[of "TYPE(32)"] sints32 by auto then have w2Less:"sint ((scast w2)::64 Word.word) < 9223372039002259455" by auto have w2Range: "-(2 ^ (size ((scast w2)::64 Word.word) - 1)) \ sint ((scast w2)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ sint ((scast w2)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w2)::64 Word.word) - 1) - 1" apply(auto simp add: Word.word_size scast_eq1 upcast_max sints64 max_less) using max_diff_pos max_less w2Less w2Geq by auto have w2case1a:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF w2Range]) have w1Lower:"sint ((scast w1)::64 Word.word) \ - (2 ^ 31) " using Word.word_sint.Rep[of "(w1)::32 Word.word"] sints32 Word.word_size scast_eq1 scast_eq2 scast_eq3 len32 mem_Collect_eq by auto have w1Leq:"sint ((scast w1)::64 Word.word) \ 2 ^ 31" apply (auto simp add: Word.word_sint.Rep[of "(w1)::32 Word.word"] sints32 scast_eq1 len32) using Word.word_sint.Rep[of "(w1)::32 Word.word"] len32[of "TYPE(32)"] sints32 by auto then have w1Less:"sint ((scast w1)::64 Word.word) < 9223372039002259455" using max_less by auto have w1MinusBound:" - (2 ^ (size ((scast w1)::64 Word.word) - 1)) \ sint ((scast w1)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ sint ((scast w1)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w1)::64 Word.word) - 1) - 1" apply(auto simp add: Word.word_size[of "(scast w1)::64 Word.word"] Word.word_size[of "(scast (-0x7FFFFFFF))::64 Word.word"] scast_eq3 scast_eq2 Word.word_sint.Rep[of "(w1)::32 Word.word"] Word.word_sint.Rep[of "0x80000001::32 Word.word"] Word.word_sint.Rep[of "(scast w1)::64 Word.word"] Word.word_sint.Rep[of "-0x7FFFFFFF::64 Word.word"] sints64 sints32) using w1Lower w1Less by auto have w1case1a:"sint (((scast w1)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w1)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by (rule signed_arith_sint(1)[of "(scast w1)::64 Word.word" "(- 0x7FFFFFFF)", OF w1MinusBound]) have w1case1a':"sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) = sint ((scast w1)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word)" using min_extend_val w1case1a by auto have w1Leq':"sint w1 \ 2^31 - 1" using Word.word_sint.Rep[of "(w1)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"]) have neg64:"(((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) = ((scast w2)::64 Word.word) + (-0x7FFFFFFF)" by auto have arith:"\x::int. x \ 2 ^ 31 - 1 \ x + (- 2147483647) < 2147483647" by auto obtain r'\<^sub>1 and r'\<^sub>2 where geq1:"r'\<^sub>1\r1" and equiv1:"w1 \\<^sub>E r'\<^sub>1" and geq2:"r'\<^sub>2\r2" and equiv2:"w2 \\<^sub>E r'\<^sub>2" using up1 up2 unfolding repU_def by auto show ?thesis proof (cases rule: case_pu_inf[where ?w1.0=w1, where ?w2.0=w2]) case PosAny then show ?thesis apply (auto simp add: repU_def repe.simps) using linear by blast next case AnyPos then show ?thesis apply (auto simp add: repU_def repe.simps) using linear by blast next case NegNeg then show ?thesis using up1 up2 by (auto simp add: repU_def repe.simps) next case NegNum assume neq1:"w2 \ POS_INF" assume eq2:"w1 = NEG_INF" assume neq3:"w2 \ NEG_INF" let ?sum = "(scast w2 + scast NEG_INF)::64 Word.word" have leq1:"r'\<^sub>1 \ (real_of_int (sint NEG_INF))" using equiv1 neq1 eq2 neq3 by (auto simp add: repe.simps) have leq2:"r'\<^sub>2 = (real_of_int (sint w2))" using equiv2 neq1 eq2 neq3 by (auto simp add: repe.simps) have case1:"?sum <=s ((scast NEG_INF)::64 Word.word) \ NEG_INF \\<^sub>U r1 + r2" using up1 up2 apply (simp add: repU_def repe.simps word_sle_def) apply(rule exI[where x= "r1 + r2"]) apply(auto) using w2case1a min_extend_neg apply (auto simp add: neq1 eq2 neq3 repINT repU_def repe.simps repeInt_simps up2 word_sless_alt) using repINT repU_def repe.simps repeInt_simps up2 word_sless_alt add.right_neutral add_mono dual_order.trans of_int_le_0_iff scast_eq3 by fastforce+ have case2:"\(?sum <=s scast NEG_INF) \ scast ?sum \\<^sub>U r1 + r2" apply(simp add: repU_def repe.simps word_sle_def up1 up2) apply(rule exI[where x= "r'\<^sub>2 - 0x7FFFFFFF"]) apply(rule conjI) subgoal proof - assume " \ sint (scast w2 + 0xFFFFFFFF80000001) \ - 2147483647" have bound1:"r1 \ - 2147483647" using leq1 geq1 by (auto) have bound2:"r2 \ r'\<^sub>2" using leq2 geq2 by auto show "r1 + r2 \ r'\<^sub>2 - 2147483647" using bound1 bound2 by(linarith) qed apply(rule disjI2) apply(rule disjI2) apply(auto) subgoal proof - assume a:"\ sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) \ - 2147483647" then have sintw2_bound:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding min_extend_val by auto have case1a:" sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF w2Range]) have "- 0x7FFFFFFF < sint w2 + (- 0x7FFFFFFF)" using sintw2_bound case1a min_extend_val' scast_eq3 by linarith then have w2bound:"0 < sint w2" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w2)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" using case1a scast_eq3 min_extend_val' Word.word_sint.Rep[of "(w2)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF)::64 Word.word)) " using scast_down_range[OF rightSize] by auto then have b:"sint ((scast (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001))::word) = sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001)" using min_extend_val by auto have c:"sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) = sint ((scast w2)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word)" using min_extend_val case1a by auto show "r'\<^sub>2 - 2147483647 = (real_of_int (sint ((scast (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001))::word)))" using a b min_extend_val' scast_eq3 leq2 case1a by auto qed subgoal proof - have range2a:" - (2 ^ (size ((scast w2)::64 Word.word) - 1)) \ sint ((scast w2)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ sint ((scast w2)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w2)::64 Word.word) - 1) - 1" apply(auto simp add: Word.word_size scast_eq1 upcast_max sints64 sints32 max_less) using max_diff_pos max_less w2Geq w2Less by auto have case2a:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF range2a]) have neg64:"(((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) = ((scast w2)::64 Word.word) + (-0x7FFFFFFF)" by auto assume "\ sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) \ - 2147483647" then have sintw2_bound:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding neg64 by auto have a:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF)) = sint((scast w2)::64 Word.word) + sint((-0x7FFFFFFF)::64 Word.word)" using case2a by auto have b:"sint ((scast w2)::64 Word.word) = sint w2" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have d:"sint w2 \ 2^31 - 1" using Word.word_sint.Rep[of "(w2)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"]) have "- 0x7FFFFFFF < sint w2 + (- 0x7FFFFFFF)" using sintw2_bound case2a min_extend_val' scast_eq3 by linarith then have w2bound:"0 < sint w2" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w2)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case2a b min_extend_val' using Word.word_sint.Rep[of "(w2)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto have "sint (scast (((scast w2)::64 Word.word) + (-0x7FFFFFFF))::word) < 2147483647" unfolding downcast a b min_extend_val' using range2s[of "sint w2", OF d] by auto then show "sint (scast (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001)::word) < 2147483647" by auto qed subgoal proof - assume notLeq:"\ sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) \ - 2147483647" then have gr:"sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) > - 2147483647" by auto have case2a:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF w2Range]) from neg64 have sintw2_bound:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding neg64 using notLeq by auto have a:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF)) = sint((scast w2)::64 Word.word) + sint((-0x7FFFFFFF)::64 Word.word)" using case2a by auto have c:"sint((-0x7FFFFFFF)::64 Word.word) = -0x7FFFFFFF" by auto have d:"sint w2 \ 2^31 - 1" using Word.word_sint.Rep[of "(w2)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"]) have "- 0x7FFFFFFF < sint w2 + (- 0x7FFFFFFF)" using sintw2_bound case2a c scast_eq3 by linarith then have w2bound:"0 < sint w2" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w2)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case2a scast_eq3 using Word.word_sint.Rep[of "(w2)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto have sintEq:" sint ((scast (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001))::word) = sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) " using downcast by auto show "-2147483647 < real_of_int (sint ((scast (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001))::word))" unfolding sintEq using gr of_int_less_iff of_int_minus of_int_numeral by linarith qed done have castEquiv:"\(?sum <=s scast NEG_INF) \ (scast ?sum) \\<^sub>U r1 + r2" using up1 up2 case1 case2 by fastforce have letRep:"(let sum = ?sum in if sum <=s scast NEG_INF then NEG_INF else scast sum) \\<^sub>U r1 + r2" using case1 case2 by(cases "?sum <=s scast NEG_INF"; auto) show "pu w1 w2 \\<^sub>U r1 + r2" using letRep eq2 neq1 by(auto) next case NumNeg assume neq3:"w1 \ NEG_INF" assume neq1:"w1 \ POS_INF" assume eq2:"w2 = NEG_INF" let ?sum = "(scast w1 + scast NEG_INF)::64 Word.word" have case1:"?sum <=s ((scast NEG_INF)::64 Word.word) \ NEG_INF \\<^sub>U r1 + r2" using up1 up2 apply (simp add: repU_def repe.simps word_sle_def) apply(rule exI[where x= "r1 + r2"]) apply(auto) using w1case1a min_extend_neg apply (auto simp add: neq1 eq2 neq3 repINT repU_def repe.simps repeInt_simps up2 word_sless_alt) using repINT repU_def repe.simps repeInt_simps up2 word_sless_alt proof - fix r' assume a1:"sint ((scast w1)::64 Word.word) \ 0" then have h3:"sint w1 \ 0" using scast_eq1 by auto assume a2:"r2 \ r'" assume a3:"r1 \ (real_of_int (sint w1))" assume a4:"r' \ (- 2147483647)" from a2 a4 have h1:"r2 \ - 2147483647" by auto from a1 a3 h3 have h2:"r1 \ 0" using dual_order.trans of_int_le_0_iff by blast show "r1 + r2 \ (- 2147483647)" using h1 h2 add.right_neutral add_mono by fastforce qed have leq1:"r'\<^sub>2 \ (real_of_int (sint NEG_INF))" and leq2:"r'\<^sub>1 = (real_of_int (sint w1))" using equiv1 equiv2 neq1 eq2 neq3 unfolding repe.simps by auto have case1a:"sint (((scast w1)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w1)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF w1MinusBound]) have case2:"\(?sum <=s scast NEG_INF) \ scast ?sum \\<^sub>U r1 + r2" apply (simp add: repU_def repe.simps word_sle_def up1 up2) apply(rule exI[where x= "r'\<^sub>1 - 0x7FFFFFFF"]) (*r1 + r2*) apply(rule conjI) subgoal using leq1 leq2 geq1 geq2 by auto apply(rule disjI2) apply(rule disjI2) apply(auto) subgoal proof - have f:"r'\<^sub>1 = (real_of_int (sint w1))" by (simp add: leq1 leq2 ) assume a:"\ sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) \ - 2147483647" then have sintw2_bound:"sint (((scast w1)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding min_extend_val by auto have "- 0x7FFFFFFF < sint w1 + (- 0x7FFFFFFF)" using sintw2_bound case1a min_extend_val' scast_eq1 by linarith then have w2bound:"0 < sint w1" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding w1case1a using w2bound Word.word_sint.Rep[of "(w1)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"] scast_eq1) have downcast:"sint ((scast (((scast w1)::64 Word.word) + ((- 0x7FFFFFFF))))::word) = sint (((scast w1)::64 Word.word) + ((- 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto then have b:"sint ((scast (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001))::word) = sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001)" using min_extend_val by auto have "(real_of_int (sint ((scast w1)::64 Word.word ) - 2147483647)) = r'\<^sub>1 - (real_of_int 2147483647)" by (simp add: scast_eq1 leq2) then show "r'\<^sub>1 - 2147483647 = (real_of_int (sint ((scast (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001))::word)))" by (metis b w1case1a' min_extend_val' diff_minus_eq_add minus_minus of_int_numeral) qed subgoal proof - assume "\ sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) \ - 2147483647" then have sintw2_bound:"sint (((scast w1)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding neg64 by auto have "- 0x7FFFFFFF < sint w1 + (- 0x7FFFFFFF)" using sintw2_bound case1a min_extend_val' scast_eq1 by linarith then have w2bound:"0 < sint w1" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a scast_eq1 w1case1a' using Word.word_sint.Rep[of "(w1)::32 Word.word"] w2bound by(auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w1)::64 Word.word) + ((- 0x7FFFFFFF))))::word) = sint (((scast w1)::64 Word.word) + ((- 0x7FFFFFFF)::64 Word.word)) " using scast_down_range[OF rightSize] by auto show "sint (scast (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001)::word) < 2147483647" using downcast min_extend_val' w1case1a' scast_eq1 arith[of "sint w1", OF w1Leq'] by auto qed subgoal proof - assume notLeq:"\ sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) \ - 2147483647" then have gr:"sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) > - 2147483647" by auto then have sintw2_bound:"sint (((scast w1)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding neg64 using notLeq by auto have "- 0x7FFFFFFF < sint w1 + (- 0x7FFFFFFF)" using sintw2_bound case1a min_extend_val' scast_eq1 by linarith then have w2bound:"0 < sint w1" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a scast_eq1 w1case1a' using Word.word_sint.Rep[of "(w1)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) show "- 2147483647 < real_of_int (sint ((scast (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001))::word))" using scast_down_range[OF rightSize] gr of_int_less_iff of_int_minus of_int_numeral by auto qed done have letUp:"(let sum=?sum in if sum <=s scast NEG_INF then NEG_INF else scast sum) \\<^sub>U r1+r2" using case1 case2 by (auto simp add: Let_def) have puSimp:"pu w1 w2=(let sum = ?sum in if sum <=s scast NEG_INF then NEG_INF else scast sum)" using neq3 neq1 eq2 equiv1 leq2 repeInt_simps by force then show "pu w1 w2 \\<^sub>U r1 + r2" using letUp puSimp by auto next case NumNum assume notinf1:"w1 \ POS_INF" assume notinf2:"w2 \ POS_INF" assume notneginf1:"w1 \ NEG_INF" assume notneginf2:"w2 \ NEG_INF" let ?sum = "((scast w1)::64 Word.word) + ((scast w2):: 64 Word.word)" have inf_case:"scast POS_INF <=s ?sum \ POS_INF \\<^sub>U r1 + r2" using repU_def repePos_simps by (meson dual_order.strict_trans not_less order_refl) have truth:" - (2 ^ (size ((scast w1)::64 Word.word) - 1)) \ sint ((scast w1)::64 Word.word) + sint ((scast w2)::64 Word.word) \ sint ((scast w1)::64 Word.word) + sint ((scast w2)::64 Word.word) \ 2 ^ (size ((scast w1)::64 Word.word) - 1) - 1" using Word.word_size[of "(scast w2)::64 Word.word"] Word.word_size[of "(scast w1)::64 Word.word"] scast_eq1 scast_eq3 Word.word_sint.Rep[of "(w1)::32 Word.word"] Word.word_sint.Rep[of "(w2)::32 Word.word"] Word.word_sint.Rep[of "(scast w1)::64 Word.word"] Word.word_sint.Rep[of "(scast w2)::64 Word.word"] sints64 sints32 by auto have sint_eq:"sint((scast w1 + scast w2)::64 Word.word) = sint w1 + sint w2" using signed_arith_sint(1)[of "(scast w1)::64 Word.word" "(scast w2)::64 Word.word", OF truth] scast_eq1 scast_eq3 by auto have bigOne:"scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word) \ \r'\r1 + r2. r' \ (- 0x7FFFFFFF)" proof - assume "scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word)" then have sum_leq:"sint w1 + sint w2 \ - 0x7FFFFFFF" and sum_leq':" (sint w1 + sint w2) \ (- 2147483647)" using sint_eq unfolding Word.word_sle_def by auto obtain r'\<^sub>1 r'\<^sub>2 ::real where bound1:"r'\<^sub>1 \ r1 \ (w1 \\<^sub>E r'\<^sub>1)" and bound2:"r'\<^sub>2 \ r2 \ (w2 \\<^sub>E r'\<^sub>2)" using up1 up2 unfolding repU_def by auto have somethingA:"r'\<^sub>1 \ sint w1" and somethingB:"r'\<^sub>2 \ sint w2" using \scast w1 + scast w2 <=s - 0x7FFFFFFF\ word_sle_def notinf1 notinf2 bound1 bound2 unfolding repe.simps by auto have something:"r'\<^sub>1 + r'\<^sub>2 \ sint w1 + sint w2" using somethingA somethingB add_mono by fastforce show "\r'\r1 + r2. r' \ (- 0x7FFFFFFF)" apply(rule exI[where x = "r'\<^sub>1 + r'\<^sub>2"]) using bound1 bound2 add_mono something sum_leq' order.trans by auto qed have anImp:"\r'. (r'\r1 + r2 \ r' \ (- 2147483647)) \ (\r. - (2 ^ 31 - 1) = - (2 ^ 31 - 1) \ r' = r \ r \ (real_of_int (sint ((- (2 ^ 31 - 1))::32 Word.word))))" by auto have anEq:"((scast ((- (2 ^ 31 - 1))::32 Word.word))::64 Word.word) = (- 0x7FFFFFFF)" by auto have bigTwo: "\(((scast POS_INF)::64 Word.word) <=s ?sum) \ \(?sum <=s ((scast NEG_INF)::64 Word.word)) \ \r'\r1 + r2. r' = (real_of_int (sint (scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word))::word))) \ (r' < 0x7FFFFFFF \ (-0x7FFFFFFF) < r')" proof - assume "\(((scast POS_INF)::64 Word.word) <=s ?sum)" then have sum_leq:"sint w1 + sint w2 < 0x7FFFFFFF" unfolding Word.word_sle_def POS_INF_def using sint_eq by auto then have sum_leq':" (sint w1 + sint w2) < (2147483647)" by auto assume "\(?sum <=s ((scast NEG_INF)::64 Word.word))" then have sum_geq:"(- 0x7FFFFFFF) < sint w1 + sint w2" unfolding Word.word_sle_def NEG_INF_def using sint_eq by auto then have sum_geq':" (- 2147483647) < (sint w1 + sint w2)" by auto obtain r'\<^sub>1 r'\<^sub>2 ::real where bound1:"r'\<^sub>1 \ r1 \ (w1 \\<^sub>E r'\<^sub>1)" and bound2:"r'\<^sub>2 \ r2 \ (w2 \\<^sub>E r'\<^sub>2)" using up1 up2 unfolding repU_def by auto have somethingA:"r'\<^sub>1 \ sint w1" and somethingB:"r'\<^sub>2 \ sint w2" using word_sle_def notinf1 notinf2 bound1 bound2 unfolding repe.simps by auto have something:"r'\<^sub>1 + r'\<^sub>2 \ sint w1 + sint w2" using somethingA somethingB add_mono by fastforce have "(w1 \\<^sub>E r'\<^sub>1)" using bound1 by auto then have r1w1:"r'\<^sub>1 = (real_of_int (sint w1))" and w1U:" (real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and w1L:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" unfolding repe.simps using notinf1 notinf2 notneginf1 notneginf2 by (auto) have "(w2 \\<^sub>E r'\<^sub>2)" using bound2 by auto then have r2w1:"r'\<^sub>2 = (real_of_int (sint w2))" and w2U:" (real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and w2L:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" unfolding repe.simps using notinf1 notinf2 notneginf1 notneginf2 by (auto) have "sint (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)) = sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)" apply(rule scast_down_range) unfolding sint_eq using sints32 sum_geq sum_leq by auto then have cast_eq:"(sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)) = sint w1 + sint w2" using scast_down_range sints32 sum_geq sum_leq sint_eq by auto from something and cast_eq have r12_sint_scast:"r'\<^sub>1 + r'\<^sub>2 = (sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word))" using r1w1 w1U w1L r2w1 w2U w2L by (simp) show ?thesis using bound1 bound2 add_mono r12_sint_scast cast_eq sum_leq sum_leq' sum_geq' \r'\<^sub>1 + r'\<^sub>2 = (real_of_int (sint (scast (scast w1 + scast w2))))\ by auto qed have neg_inf_case:"?sum <=s ((scast ((NEG_INF)::word))::64 Word.word) \ NEG_INF \\<^sub>U r1 + r2" proof (unfold repU_def NEG_INF_def repe.simps) assume "scast w1 + scast w2 <=s ((scast ((- (2 ^ 31 - 1))::word))::64 Word.word)" then have "scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word)" by (metis anEq) then obtain r' where geq:"(r' \ r1 + r2)" and leq:"(r' \ (- 0x7FFFFFFF))" using bigOne by auto show "(\r'\plus r1 r2. (\r. uminus (minus(2 ^ 31) 1) = POS_INF \ r' = r \ (real_of_int (sint POS_INF)) \ r) \ (\r. uminus (minus(2 ^ 31) 1) = uminus (minus(2 ^ 31) 1) \ r' = r \ r \ real_of_int (sint ((uminus (minus(2 ^ 31) 1))::word))) \ (\w. uminus (minus(2 ^ 31) 1) = w \ r' = real_of_int (sint w) \ (real_of_int (sint w)) < (real_of_int (sint POS_INF)) \ less (real_of_int (sint (uminus (minus(2 ^ 31) 1)))) (real_of_int (sint w))))" using leq anImp geq by auto qed have int_case:"\(((scast POS_INF)::64 Word.word) <=s ?sum) \ \ (?sum <=s ((scast NEG_INF)::64 Word.word)) \ ((scast ?sum)::word) \\<^sub>U r1 + r2" proof - assume bound1:"\ ((scast POS_INF)::64 Word.word) <=s scast w1 + scast w2" assume bound2:"\ scast w1 + scast w2 <=s ((scast NEG_INF)::64 Word.word)" obtain r'::real where rDef:"r' = (real_of_int (sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)))" and r12:"r'\r1 + r2" and boundU:"r' < 0x7FFFFFFF" and boundL:"(-0x7FFFFFFF) < r'" using bigTwo[OF bound1 bound2] by auto obtain w::word where wdef:"w = (scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word))::word)" by auto then have wr:"r' = (real_of_int (sint w))" using r12 bound1 bound2 rDef by blast show ?thesis unfolding repU_def repe.simps using r12 wdef rDef boundU boundL wr by auto qed have almost:"(let sum::64 Word.word = scast w1 + scast w2 in if scast POS_INF <=s sum then POS_INF else if sum <=s scast NEG_INF then NEG_INF else scast sum) \\<^sub>U r1 + r2" apply(cases "((scast POS_INF)::64 Word.word) <=s ((?sum)::64 Word.word)") subgoal using inf_case Let_def int_case neg_inf_case by auto apply(cases "?sum <=s scast NEG_INF") subgoal using inf_case Let_def int_case neg_inf_case proof - assume "\ (scast POS_INF::64 Word.word) <=s scast w1 + scast w2" then have "\ (scast w1::64 Word.word) + scast w2 <=s scast NEG_INF \ \ (scast POS_INF::64 Word.word) <=s scast w1 + scast w2 \ \ (scast w1::64 Word.word) + scast w2 <=s scast NEG_INF \ ((let w = scast w1 + scast w2 in if scast POS_INF <=s (w::64 Word.word) then POS_INF else if w <=s scast NEG_INF then NEG_INF else scast w) \\<^sub>U r1 + r2)" using neg_inf_case by presburger then show ?thesis using int_case by force qed subgoal using inf_case Let_def int_case neg_inf_case proof - assume a1: "\ (scast POS_INF::64 Word.word) <=s scast w1 + scast w2" assume "\ (scast w1::64 Word.word) + scast w2 <=s scast NEG_INF" have "\ (scast w1::64 Word.word) + scast w2 <=s scast NEG_INF \ \ (scast POS_INF::64 Word.word) <=s scast w1 + scast w2 \ ((let w = scast w1 + scast w2 in if scast POS_INF <=s (w::64 Word.word) then POS_INF else if w <=s scast NEG_INF then NEG_INF else scast w) \\<^sub>U r1 + r2)" using a1 neg_inf_case by presburger then show ?thesis using int_case by force qed done then show ?thesis using notinf1 notinf2 notneginf1 notneginf2 by auto qed qed text\Lower bound of w1 + w2\ fun pl :: "word \ word \ word" where "pl w1 w2 = (if w1 = NEG_INF then NEG_INF else if w2 = NEG_INF then NEG_INF else if w1 = POS_INF then (if w2 = POS_INF then POS_INF else (let sum::64 Word.word = ((scast w2)::64 Word.word) + ((scast POS_INF)::64 Word.word) in if ((scast POS_INF)::64 Word.word) <=s(sum::64 Word.word) then POS_INF else scast sum)) else if w2 = POS_INF then (let sum::64 Word.word = ((scast w1)::64 Word.word) + ((scast POS_INF)::64 Word.word) in if ((scast POS_INF)::64 Word.word) <=s(sum::64 Word.word) then POS_INF else scast sum) else (let sum::64 Word.word = ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word) in if ((scast POS_INF)::64 Word.word) <=s (sum::64 Word.word) then POS_INF else if (sum::64 Word.word) <=s ((scast NEG_INF)::64 Word.word) then NEG_INF else scast sum))" subsection \Addition lower bound\ text\Correctness of lower bound of w1 + w2\ lemma pl_lemma: assumes lo1:"w1 \\<^sub>L (r1::real)" assumes lo2:"w2 \\<^sub>L (r2::real)" shows "pl w1 w2 \\<^sub>L (r1 + r2)" proof - have scast_eq1:"sint((scast w1)::64 Word.word) = sint w1" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have scast_eq2:"sint((scast (0x80000001::word))::64 Word.word)=sint((0x80000001::32 Word.word))" by auto have scast_eq3:"sint((scast w2)::64 Word.word) = sint w2" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have sints64:"sints 64 = {i. - (2 ^ 63) \ i \ i < 2 ^ 63}" using sints_def[of 64] range_sbintrunc[of 63] by auto have sints32:"sints 32 = {i. - (2 ^ 31) \ i \ i < 2 ^ 31}" using sints_def[of 32] range_sbintrunc[of 31] by auto have thing1:"0 \ 9223372034707292161 + ((-(2 ^ 31))::real)" by auto have "sint (( w2)) \ (-(2 ^ 31))" using Word.word_sint.Rep[of "(w2)::32 Word.word"] sints32 mem_Collect_eq Word.word_size[of "(scast w2)::64 Word.word"] scast_eq1 scast_eq2 scast_eq3 len32 by auto then have thing4:"sint ((scast w2)::64 Word.word) \ (-(2 ^ 31))" using scast_down_range sint_ge sints_num using scast_eq3 by linarith have aLeq2:"(-(2 ^ 31)::int) \ -9223372039002259455" by auto then have thing2:" (0::int) \ 9223372039002259455 + sint ((scast w2)::64 Word.word)" using thing4 aLeq2 by (metis ab_group_add_class.ab_left_minus add.commute add_mono neg_le_iff_le) have aLeq:"2 ^ 31 \ (9223372039002259455::int)" by auto have bLeq:"sint ((scast w2)::64 Word.word) \ 2 ^ 31" apply ( auto simp add: Word.word_sint.Rep[of "(w2)::32 Word.word"] sints32 scast_eq3 len32) using Word.word_sint.Rep[of "(w2)::32 Word.word"] len32[of "TYPE(32)"] sints32 by auto have thing3:" sint ((scast w2)::64 Word.word) \ 9223372034707292160 " using aLeq bLeq by auto have truth:" - (2 ^ (size ((scast w2)::64 Word.word) - 1)) \ sint ((scast w2)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ sint ((scast w2)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w2)::64 Word.word) - 1) - 1" by(auto simp add: Word.word_size[of "(scast w2)::64 Word.word"] Word.word_size[of "(scast (0x7FFFFFFF))::64 Word.word"] scast_eq1 scast_eq2 sints64 sints32 thing2 thing1 thing3) have case1a:" sint (((scast w2)::64 Word.word) + (0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF truth]) have case1b:"sint ((0xFFFFFFFF80000001)::64 Word.word) < 0" by auto have arith:"\x::int. x \ 2 ^ 31 - 1 \ x + (- 2147483647) < 2147483647" by auto have neg64:"(((scast w2)::64 Word.word) + 0x7FFFFFFF) = ((scast w2)::64 Word.word) + (0x7FFFFFFF)" by auto obtain r'\<^sub>1 and r'\<^sub>2 where geq1:"r'\<^sub>1\r1" and equiv1:"w1 \\<^sub>E r'\<^sub>1" and geq2:"r'\<^sub>2\r2" and equiv2:"w2 \\<^sub>E r'\<^sub>2" using lo1 lo2 unfolding repL_def by auto show ?thesis proof (cases rule: case_pl_inf[where ?w1.0=w1, where ?w2.0=w2]) case NegAny then show ?thesis apply (auto simp add: repL_def repe.simps) using lo1 lo2 linear by auto next case AnyNeg then show ?thesis apply (auto simp add: repL_def repe.simps) using linear by auto next case PosPos then show ?thesis using lo1 lo2 by (auto simp add: repL_def repe.simps) next case PosNum assume neq1:"w2 \ POS_INF" assume eq2:"w1 = POS_INF" assume neq3:"w2 \ NEG_INF" let ?sum = "(scast w2 + scast POS_INF)::64 Word.word" have case1:"(((scast POS_INF)::64 Word.word) <=s ?sum) \ POS_INF \\<^sub>L r1 + r2" using lo1 lo2 apply (simp add: repL_def repe.simps word_sle_def) apply(rule exI[where x= "r1 + r2"]) using case1a case1b apply (auto simp add: neq1 eq2 neq3 repINT repL_def repe.simps repeInt_simps lo2 word_sless_alt) proof - fix r' assume a1:"0 \ sint (((scast w2)::64 Word.word))" from a1 have h3:"2147483647 \ sint w2 + 0x7FFFFFFF " using scast_eq3 by auto assume a2:"r' \ r1" assume a3:"(real_of_int (sint w2)) \ r2" assume a4:"(2147483647) \ r'" from a2 a4 have h1:"2147483647 \ r1" by auto from a1 a3 h3 have h2:"0 \ r2" using dual_order.trans of_int_le_0_iff le_add_same_cancel2 by fastforce show "(2147483647) \ r1 + r2 " using h1 h2 h3 add.right_neutral add_mono by fastforce qed have leq1:"r'\<^sub>1 \ (real_of_int (sint POS_INF))" using equiv1 neq1 eq2 neq3 unfolding repe.simps by auto have leq2:"r'\<^sub>2 = (real_of_int (sint w2))" using equiv2 neq1 eq2 neq3 unfolding repe.simps by auto have case2:"\(scast POS_INF <=s ?sum) \ scast ?sum \\<^sub>L r1 + r2" apply (simp add: repL_def repe.simps word_sle_def lo1 lo2) apply(rule exI[where x= "r'\<^sub>2 + 0x7FFFFFFF"]) (*r1 + r2*) apply(rule conjI) subgoal proof - assume "\ 2147483647 \ sint (scast w2 + 0x7FFFFFFF)" have bound1:"2147483647 \ r1" using leq1 geq1 by (auto) have bound2:"r'\<^sub>2 \ r2 " using leq2 geq2 by auto show "r'\<^sub>2 + 2147483647 \ r1 + r2" using bound1 bound2 by linarith qed apply(rule disjI2) apply(rule disjI2) apply(auto) subgoal proof - assume a:"\ 2147483647 \ sint (((scast w2)::64 Word.word) + 0x7FFFFFFF)" then have sintw2_bound:"2147483647 > sint (((scast w2)::64 Word.word) + (0x7FFFFFFF))" by auto have case1a:"sint (((scast w2)::64 Word.word) + (0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF truth]) have a1:"sint (((scast w2)::64 Word.word) + (0x7FFFFFFF)) = sint((scast w2)::64 Word.word) + sint((0x7FFFFFFF)::64 Word.word)" using case1a by auto have c1:"sint((0x7FFFFFFF)::64 Word.word) = 0x7FFFFFFF" by auto have "sint w2 + ( 0x7FFFFFFF) < 0x7FFFFFFF" using sintw2_bound case1a c1 scast_eq3 by linarith then have w2bound:"sint w2 < 0" using add_less_same_cancel2 by blast have rightSize:"sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a scast_eq3 c1 using Word.word_sint.Rep[of "(w2)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w2)::64 Word.word) + (( 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + (( 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto then have b:"sint ((scast (((scast w2)::64 Word.word) + 0x7FFFFFFF))::word) = sint (((scast w2)::64 Word.word) + 0x7FFFFFFF)" by auto have c:"sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) = sint ((scast w2)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word)" using case1a by auto have d:"sint ((0x7FFFFFFF)::64 Word.word) = (0x7FFFFFFF)" by auto have f:"r'\<^sub>2 = (real_of_int (sint w2))" by (simp add: leq2) show "r'\<^sub>2 + 2147483647 = (real_of_int (sint ((scast (((scast w2)::64 Word.word) + 0x7FFFFFFF))::word)))" using a b c d scast_eq3 f leq2 of_int_numeral by fastforce qed subgoal proof - have truth2a:"-(2^(size ((scast w2)::64 Word.word)-1)) \ sint ((scast w2)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ sint ((scast w2)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w2)::64 Word.word) - 1) - 1" apply(auto simp add: Word.word_size[of "(scast w2)::64 Word.word"] Word.word_size[of "(scast (0x7FFFFFFF))::64 Word.word"] scast_eq1 scast_eq2 sints64 sints32 thing2) using thing1 thing2 thing3 by auto have case2a:" sint (((scast w2)::64 Word.word) + (0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF truth2a]) have min_cast:"(0x7FFFFFFF::64 Word.word) =((scast (0x7FFFFFFF::word))::64 Word.word)" by auto assume "\ 2147483647 \ sint (((scast w2)::64 Word.word) + 0x7FFFFFFF)" then have sintw2_bound:"2147483647 > sint (((scast w2)::64 Word.word) + (0x7FFFFFFF))" using neg64 by auto have a:"sint (((scast w2)::64 Word.word) + (0x7FFFFFFF)) = sint((scast w2)::64 Word.word) + sint((0x7FFFFFFF)::64 Word.word)" using case2a by auto have c:"sint((0x7FFFFFFF)::64 Word.word) = 0x7FFFFFFF" by auto have " 0x7FFFFFFF > sint w2 + ( 0x7FFFFFFF)" using sintw2_bound case2a c scast_eq3 by linarith then have w2bound:" sint w2 < 0" by simp have rightSize:"sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case2a scast_eq3 c apply (auto simp add: sints32 len32[of "TYPE(32)"]) using Word.word_sint.Rep[of "(w2)::32 Word.word"] sints32 len32[of "TYPE(32)"] w2bound by auto have downcast:"sint ((scast (((scast w2)::64 Word.word) + (( 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + (( 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto then show "sint (scast (((scast w2)::64 Word.word) + 0x7FFFFFFF)::word) < 2147483647" unfolding downcast a scast_eq3 c using w2bound by auto qed subgoal proof - assume notLeq:"\ 2147483647 \ sint (((scast w2)::64 Word.word) + 0x7FFFFFFF)" then have gr:"sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) < 2147483647" by auto have case2a:" sint (((scast w2)::64 Word.word) + (0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF truth]) have min_cast:"(0x7FFFFFFF::64 Word.word) =((scast (0x7FFFFFFF::word))::64 Word.word)" by auto have neg64:"(((scast w2)::64 Word.word) + 0x7FFFFFFF) = ((scast w2)::64 Word.word) + (0x7FFFFFFF)" by auto then have sintw2_bound:"sint (((scast w2)::64 Word.word) + (0x7FFFFFFF)) < 2147483647" using neg64 using notLeq by auto have a:"sint (((scast w2)::64 Word.word) + (0x7FFFFFFF)) = sint((scast w2)::64 Word.word) + sint((0x7FFFFFFF)::64 Word.word)" using case2a by auto have c:"sint((0x7FFFFFFF)::64 Word.word) = 0x7FFFFFFF" by auto have "- 2147483647 \ w2" using neq3 unfolding NEG_INF_def by auto then have "sint((- 2147483647)::word) \ sint w2" using word_sint.Rep_inject by blast then have n1:"- 2147483647 \ sint w2" by auto have "- 2147483648 \ w2" apply(rule repe.cases[OF equiv2]) by auto then have "sint(- 2147483648::word) \ sint w2" using word_sint.Rep_inject by blast then have n2:"- 2147483648 \ sint w2" by auto then have d:"sint w2 > - 2147483647" using Word.word_sint.Rep[of "(w2)::32 Word.word"] sints32 len32[of "TYPE(32)"] neq3 n1 n2 by auto have w2bound:"- 2147483647 < sint w2 + 0x7FFFFFFF" using sintw2_bound case2a c scast_eq3 d by linarith have rightSize:"sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" using sints32 len32[of "TYPE(32)"] w2bound Word.word_sint.Rep[of "(w2)::32 Word.word"] c case2a scast_eq3 sintw2_bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w2)::64 Word.word) + (( 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + (( 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto have sintEq:" sint ((scast (((scast w2)::64 Word.word) + 0x7FFFFFFF))::word) = sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) " using downcast by auto show "- 2147483647 < real_of_int (sint ((scast (((scast w2)::64 Word.word) + 0x7FFFFFFF))::word))" unfolding sintEq using gr of_int_less_iff of_int_minus of_int_numeral c case2a scast_eq3 w2bound by linarith qed done have "(let sum = ?sum in if scast POS_INF <=s sum then POS_INF else scast sum) \\<^sub>L r1 + r2" using case1 case2 by (auto simp add: Let_def) then show ?thesis using lo1 lo2 neq1 eq2 neq3 by (auto) next case NumPos assume neq3:"w1 \ NEG_INF" assume neq1:"w1 \ POS_INF" assume eq2:"w2 = POS_INF" let ?sum = "(scast w1 + scast POS_INF)::64 Word.word" have thing1:"0 \ 9223372034707292161 + ((-(2 ^ 31))::real)" by auto have "sint (( w1)) \ (-(2 ^ 31))" using Word.word_sint.Rep[of "(w1)::32 Word.word"] scast_eq1 scast_eq2 scast_eq3 Word.word_size[of "(scast w1)::64 Word.word"] sints32 len32 mem_Collect_eq by auto then have thing4:"sint ((scast w1)::64 Word.word) \ (-(2 ^ 31))" using scast_down_range sint_ge sints_num scast_eq3 scast_eq1 by linarith have aLeq2:"(-(2 ^ 31)::int) \ -9223372039002259455" by auto then have thing2:" (0::int) \ 9223372039002259455 + sint ((scast w1)::64 Word.word)" using thing4 aLeq2 by (metis ab_group_add_class.ab_left_minus add.commute add_mono neg_le_iff_le) have aLeq:"2 ^ 31 \ (9223372039002259455::int)" by auto have bLeq:"sint ((scast w1)::64 Word.word) \ 2 ^ 31" apply (auto simp add: Word.word_sint.Rep[of "(w1)::32 Word.word"] sints32 scast_eq1 len32) using Word.word_sint.Rep[of "(w1)::32 Word.word"] len32[of "TYPE(32)"] sints32 by clarsimp have thing3:" sint ((scast w1)::64 Word.word) \ 9223372034707292160 " using aLeq bLeq by auto have truth:" - (2 ^ (size ((scast w1)::64 Word.word) - 1)) \ sint ((scast w1)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ sint ((scast w1)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w1)::64 Word.word) - 1) - 1" by(auto simp add: Word.word_size[of "(scast w1)::64 Word.word"] Word.word_size[of "(scast (0x7FFFFFFF))::64 Word.word"] scast_eq3 scast_eq2 sints64 sints32 thing2 thing1 thing3) have case1a:"sint (((scast w1)::64 Word.word) + (0x7FFFFFFF::64 Word.word)) = sint ((scast w1)::64 Word.word) + sint (0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF truth]) have case1b:"sint ((0xFFFFFFFF80000001)::64 Word.word) < 0" by auto have g:"(0x7FFFFFFF::64 Word.word) = 0x7FFFFFFF" by auto have c:"sint (((scast w1)::64 Word.word) + 0x7FFFFFFF) = sint ((scast w1)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word)" using g case1a by blast have d:"sint ((0x7FFFFFFF)::64 Word.word) = (0x7FFFFFFF)" by auto have e:"sint ((scast w1)::64 Word.word) = sint w1" using scast_eq1 by blast have d2:"sint w1 \ 2^31 - 1" using Word.word_sint.Rep[of "(w1)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"]) have case1:"scast POS_INF <=s ?sum \ POS_INF \\<^sub>L r1 + r2" using lo1 lo2 apply (simp add: repL_def repe.simps word_sle_def) apply(rule exI[where x= "r1 + r2"]) apply(auto) using case1a case1b apply (auto simp add: neq1 eq2 neq3 repINT repL_def repe.simps repeInt_simps lo2 word_sless_alt) (* close 4 goals *) proof - fix r' have h3:"sint (((scast w1)::64 Word.word) + 0x7FFFFFFF) = sint (((scast w1)::64 Word.word)) + sint(0x7FFFFFFF::64 Word.word)" using case1a by auto have h4:"sint(0x7FFFFFFF::64 Word.word) = 2147483647" by auto assume a1:"0 \ sint ((scast w1)::64 Word.word)" then have h3:"sint w1 \ 0" using scast_eq1 h3 h4 by auto assume a2:"r' \ r2" assume a3:"(real_of_int (sint w1)) \ r1" assume a4:"(2147483647) \ r'" from a2 a4 have h1:"r2 \ 2147483647" by auto from a1 a3 h3 have h2:"r1 \ 0" using dual_order.trans of_int_0_le_iff by fastforce show " 2147483647 \ r1 + r2" using h1 h2 add.right_neutral add_mono by fastforce qed have leq1:"r'\<^sub>2 \ (real_of_int (sint POS_INF))" and leq2:"r'\<^sub>1 = (real_of_int (sint w1))" using equiv1 equiv2 neq1 eq2 neq3 unfolding repe.simps by auto have neg64:"(((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) = ((scast w1)::64 Word.word) + (-0x7FFFFFFF)" by auto have case2:"\(scast POS_INF <=s ?sum) \ scast ?sum \\<^sub>L r1 + r2" apply (simp add: repL_def repe.simps word_sle_def lo1 lo2) apply(rule exI[where x= "r'\<^sub>1 + 0x7FFFFFFF"]) apply(rule conjI) subgoal proof - assume "\ 2147483647 \ sint (scast w1 + 0x7FFFFFFF)" have bound1:"r2 \ 2147483647" using leq1 geq2 by (auto) have bound2:"r1 \ r'\<^sub>1" using leq2 geq1 by auto show "r'\<^sub>1 + 2147483647 \ r1 + r2" using bound1 bound2 by linarith qed apply(rule disjI2) apply(rule disjI2) apply(auto) subgoal proof - have f:"r'\<^sub>1 = (real_of_int (sint w1))" by (simp add: leq1 leq2 ) assume a:"\ 2147483647 \ sint (((scast w1)::64 Word.word) + 0x7FFFFFFF)" then have sintw2_bound:"2147483647 > sint (((scast w1)::64 Word.word) + (0x7FFFFFFF))" by auto have "0x7FFFFFFF > sint w1 + (0x7FFFFFFF)" using sintw2_bound case1a d scast_eq1 by linarith then have w2bound:"0 > sint w1" using add_less_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a e using w2bound Word.word_sint.Rep[of "(w1)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"] ) have arith:"\x::int. x \ 2 ^ 31 - 1 \ x + (- 2147483647) < 2147483647" by auto have downcast:"sint ((scast (((scast w1)::64 Word.word) + (( 0x7FFFFFFF))))::word) = sint (((scast w1)::64 Word.word) + (( 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto then have b:"sint((scast (((scast w1)::64 Word.word) + 0x7FFFFFFF))::word) = sint(((scast w1)::64 Word.word) + 0x7FFFFFFF)" using g by auto show "r'\<^sub>1 + 2147483647 = (real_of_int (sint ((scast (((scast w1)::64 Word.word) + 0x7FFFFFFF))::word)))" using a b c d e f proof - have "(real_of_int (sint ((scast w1)::64 Word.word ) + 2147483647)) = r'\<^sub>1 + (real_of_int 2147483647)" using e leq2 by auto then show ?thesis using b c d of_int_numeral by metis qed qed subgoal proof - assume "\ 2147483647 \ sint (((scast w1)::64 Word.word) + 0x7FFFFFFF)" then have sintw2_bound:"sint (((scast w1)::64 Word.word) + (0x7FFFFFFF)) < 2147483647" unfolding neg64 by auto have "0x7FFFFFFF > sint w1 + (0x7FFFFFFF)" using sintw2_bound case1a d scast_eq1 by linarith then have w2bound:"0 > sint w1" using add_less_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a e c using Word.word_sint.Rep[of "(w1)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have arith:"\x::int. x \ 2 ^ 31 - 1 \ x + (- 2147483647) < 2147483647" by auto have downcast:"sint ((scast (((scast w1)::64 Word.word) + 0x7FFFFFFF))::word) = sint (((scast w1)::64 Word.word) + ((0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto show "sint (scast (((scast w1)::64 Word.word) + 0x7FFFFFFF)::word) < 2147483647" using downcast d e c arith[of "sint w1", OF d2] sintw2_bound by linarith qed subgoal proof - assume notLeq:"\ 2147483647 \ sint (((scast w1)::64 Word.word) + 0x7FFFFFFF)" then have gr:"2147483647 > sint (((scast w1)::64 Word.word) + 0x7FFFFFFF)" by auto then have sintw2_bound:"sint (((scast w1)::64 Word.word) + (0x7FFFFFFF)) < 2147483647" unfolding neg64 using notLeq by auto have "0x7FFFFFFF > sint w1 + ( 0x7FFFFFFF)" using sintw2_bound case1a d scast_eq1 by linarith then have useful:"0 > sint w1" using add_less_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a e using Word.word_sint.Rep[of "(w1)::32 Word.word"] sints32 len32[of "TYPE(32)"] useful by auto have "- 2147483647 \ w1" using neq3 unfolding NEG_INF_def by auto then have "sint((- 2147483647)::word) \ sint w1" using word_sint.Rep_inject by blast then have n1:"- 2147483647 \ sint w1" by auto have "- 2147483648 \ w1" apply(rule repe.cases[OF equiv1]) using int_not_undef[of w1] by auto then have "sint(- 2147483648::word) \ sint w1" using word_sint.Rep_inject by blast then have n2:"- 2147483648 \ sint w1" by auto then have d:"sint w1 > - 2147483647" using Word.word_sint.Rep[of "(w1)::32 Word.word"] sints32 len32[of "TYPE(32)"] n1 n2 neq3 by (simp) have d2:"sint (0x7FFFFFFF::64 Word.word) > 0" by auto from d d2 have d3:"- 2147483647 < sint w1 + sint (0x7FFFFFFF::64 Word.word)" by auto have d4:"sint ((scast (((scast w1)::64 Word.word) + 0x7FFFFFFF))::word) = sint w1 + sint (0x7FFFFFFF::64 Word.word)" using case1a rightSize scast_down_range scast_eq1 by fastforce then show "-2147483647 < real_of_int (sint ((scast (((scast w1)::64 Word.word) + 0x7FFFFFFF))::word))" using d3 d4 by auto qed done have "(let sum = ?sum in if scast POS_INF <=s sum then POS_INF else scast sum) \\<^sub>L r1 + r2" using case1 case2 by (auto simp add: Let_def) then show ?thesis using neq1 eq2 neq3 by (auto) next case NumNum assume notinf1:"w1 \ POS_INF" assume notinf2:"w2 \ POS_INF" assume notneginf1:"w1 \ NEG_INF" assume notneginf2:"w2 \ NEG_INF" let ?sum = "((scast w1)::64 Word.word) + ((scast w2):: 64 Word.word)" have truth:" - (2 ^ (size ((scast w1)::64 Word.word) - 1)) \ sint ((scast w1)::64 Word.word) + sint ((scast w2)::64 Word.word) \ sint ((scast w1)::64 Word.word) + sint ((scast w2)::64 Word.word) \ 2 ^ (size ((scast w1)::64 Word.word) - 1) - 1" using Word.word_size[of "(scast w2)::64 Word.word"] Word.word_size[of "(scast w1)::64 Word.word"] scast_eq1 scast_eq3 sints64 sints32 Word.word_sint.Rep[of "(w1)::32 Word.word"] Word.word_sint.Rep[of "(w2)::32 Word.word"] by auto have sint_eq:"sint((scast w1 + scast w2)::64 Word.word) = sint w1 + sint w2" using signed_arith_sint(1)[of "(scast w1)::64 Word.word" "(scast w2)::64 Word.word", OF truth] scast_eq1 scast_eq3 by auto have bigOne:"scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word) \ \r'\r1 + r2. r' \ -0x7FFFFFFF" proof - assume "scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word)" then have sum_leq:"sint w1 + sint w2 \ - 0x7FFFFFFF" and sum_leq':" (sint w1 + sint w2) \ (- 2147483647)" using sint_eq unfolding Word.word_sle_def by auto obtain r'\<^sub>1 r'\<^sub>2 ::real where bound1:"r'\<^sub>1 \ r1 \ (w1 \\<^sub>E r'\<^sub>1)" and bound2:"r'\<^sub>2 \ r2 \ (w2 \\<^sub>E r'\<^sub>2)" using lo1 lo2 unfolding repL_def by auto have somethingA:"r'\<^sub>1 \ sint w1" and somethingB:"r'\<^sub>2 \ sint w2" using bound1 bound2 \scast w1 + scast w2 <=s -0x7FFFFFFF\ word_sle_def notinf1 notinf2 unfolding repe.simps by auto have something:"r'\<^sub>1 + r'\<^sub>2 \ sint w1 + sint w2" using somethingA somethingB add_mono by fastforce show "\r'\r1 + r2. r' \ (-0x7FFFFFFF)" apply(rule exI[where x = "r'\<^sub>1 + r'\<^sub>2"]) using bound1 bound2 add_mono something sum_leq' order.trans by auto qed have anImp:"\r'. (r'\r1 + r2 \ r' \ (- 2147483647)) \ (\r. - (2 ^ 31 - 1) = - (2 ^ 31 - 1) \ r' = r \ r \ (real_of_int (sint ((- (2 ^ 31 - 1))::32 Word.word))))" by auto have anEq:"((scast ((- (2 ^ 31 - 1))::32 Word.word))::64 Word.word) = (- 0x7FFFFFFF)" by auto have bigTwo: "\(((scast POS_INF)::64 Word.word) <=s ?sum) \ \(?sum <=s ((scast NEG_INF)::64 Word.word)) \ \r'\r1 + r2. r' = (real_of_int (sint (scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word))::word))) \ (r' < 0x7FFFFFFF \ (-0x7FFFFFFF) < r')" proof - assume "\(((scast POS_INF)::64 Word.word) <=s ?sum)" then have sum_leq:"sint w1 + sint w2 < 0x7FFFFFFF" unfolding Word.word_sle_def using sint_eq by auto then have sum_leq':" (sint w1 + sint w2) < (2147483647)" by auto assume "\(?sum <=s ((scast NEG_INF)::64 Word.word))" then have sum_geq:"(- 0x7FFFFFFF) < sint w1 + sint w2" unfolding Word.word_sle_def using sint_eq by auto then have sum_geq':" (- 2147483647) < (sint w1 + sint w2)" by auto obtain r'\<^sub>1 r'\<^sub>2 ::real where bound1:"r'\<^sub>1 \ r1 \ (w1 \\<^sub>E r'\<^sub>1)" and bound2:"r'\<^sub>2 \ r2 \ (w2 \\<^sub>E r'\<^sub>2)" using lo1 lo2 unfolding repL_def by auto have somethingA:"r'\<^sub>1 \ sint w1" and somethingB:"r'\<^sub>2 \ sint w2" using word_sle_def notinf1 notinf2 bound1 bound2 unfolding repe.simps by auto have something:"r'\<^sub>1 + r'\<^sub>2 \ sint w1 + sint w2" using somethingA somethingB add_mono by fastforce have "(w1 \\<^sub>E r'\<^sub>1)" using bound1 by auto then have r1w1:"r'\<^sub>1 = (real_of_int (sint w1))" and w1U:"(real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and w1L:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" unfolding repe.simps using notinf1 notinf2 notneginf1 notneginf2 by (auto) have "(w2 \\<^sub>E r'\<^sub>2)" using bound2 by auto then have r2w1:"r'\<^sub>2 = (real_of_int (sint w2))" and w2U:"(real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and w2L:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" unfolding repe.simps using notinf1 notinf2 notneginf1 notneginf2 by (auto) have "sint (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)) = sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)" apply(rule scast_down_range) unfolding sint_eq using sints32 sum_geq sum_leq by auto then have cast_eq:"(sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)) = sint w1 + sint w2" using scast_down_range sints32 sum_geq sum_leq sint_eq by auto from something and cast_eq have r12_sint_scast:"r'\<^sub>1 + r'\<^sub>2 = (sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word))" using r1w1 w1U w1L r2w1 w2U w2L by (simp) have leq_ref:"\x y ::real. x = y ==> x \ y" by auto show ?thesis apply(rule exI[where x="r'\<^sub>1 + r'\<^sub>2"]) apply(rule conjI) subgoal using r12_sint_scast cast_eq leq_ref r2w1 r1w1 add_mono[of r'\<^sub>1 r1 r'\<^sub>2 r2] bound1 bound2 by auto using bound1 bound2 add_mono r12_sint_scast cast_eq sum_leq sum_leq' sum_geq' sum_geq \r'\<^sub>1 + r'\<^sub>2 = (real_of_int (sint (scast (scast w1 + scast w2))))\ by auto qed have neg_inf_case:"?sum <=s ((scast ((NEG_INF)::word))::64 Word.word) \ NEG_INF \\<^sub>L r1 + r2" proof (unfold repL_def NEG_INF_def repe.simps) assume "scast w1 + scast w2 <=s ((scast ((- (2 ^ 31 - 1))::word))::64 Word.word)" then have "scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word)" by (metis anEq) then obtain r' where geq:"(r' \ r1 + r2)" and leq:"(r' \ (-0x7FFFFFFF))" using bigOne by auto show "(\r'\plus r1 r2. (\r. uminus (minus(2 ^ 31) 1) = POS_INF \ r' = r \ (real_of_int (sint POS_INF)) \ r) \ (\r. uminus (minus(2 ^ 31) 1) = uminus (minus(2 ^ 31) 1) \ r' = r \ r \ (real_of_int (sint ((uminus (minus(2 ^ 31) 1))::word)))) \ (\w. uminus (minus(2 ^ 31) 1) = w \ r' = (real_of_int (sint w)) \ (real_of_int (sint w)) < (real_of_int (sint POS_INF)) \ less ( (real_of_int (sint (uminus (minus(2 ^ 31) 1))))) ((real_of_int (sint w)))))" using leq geq by auto qed have bigThree:"0x7FFFFFFF <=s ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word) \ \r'\r1 + r2. 2147483647 \ r'" proof - assume "0x7FFFFFFF <=s ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)" then have sum_leq:"0x7FFFFFFF \ sint w1 + sint w2 " and sum_leq':" 2147483647 \ (sint w1 + sint w2)" using sint_eq unfolding Word.word_sle_def by auto obtain r'\<^sub>1 r'\<^sub>2 ::real where bound1:"r'\<^sub>1 \ r1 \ (w1 \\<^sub>E r'\<^sub>1)" and bound2:"r'\<^sub>2 \ r2 \ (w2 \\<^sub>E r'\<^sub>2)" using lo1 lo2 unfolding repL_def by auto have somethingA:"r'\<^sub>1 \ sint w1" and somethingB:"r'\<^sub>2 \ sint w2" using \ 0x7FFFFFFF <=s scast w1 + scast w2 \ word_sle_def notinf1 notinf2 bound1 bound2 repe.simps by auto have something:"r'\<^sub>1 + r'\<^sub>2 \ sint w1 + sint w2" using somethingA somethingB add_mono of_int_add by fastforce show "\r'\ r1 + r2. (2147483647) \ r'" apply(rule exI[where x = "r'\<^sub>1 + r'\<^sub>2"]) using bound1 bound2 add_mono something sum_leq' order.trans proof - have f1: " (real_of_int (sint w2)) = r'\<^sub>2" by (metis bound2 notinf2 notneginf2 repe.cases) have " (real_of_int (sint w1)) = r'\<^sub>1" by (metis bound1 notinf1 notneginf1 repe.cases) then have f2: " (real_of_int 2147483647) \ r'\<^sub>2 + r'\<^sub>1" using f1 sum_leq' by force have "r'\<^sub>2 + r'\<^sub>1 \ r2 + r1" by (meson add_left_mono add_right_mono bound1 bound2 order.trans) then show "r'\<^sub>1 + r'\<^sub>2 \ r1 + r2 \ 2147483647 \ r'\<^sub>1 + r'\<^sub>2" using f2 by (simp add: add.commute) qed qed have inf_case:"((scast POS_INF)::64 Word.word) <=s ?sum \ POS_INF \\<^sub>L r1 + r2" proof - assume "((scast POS_INF)::64 Word.word) <=s ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)" then have "((scast ((2 ^ 31 - 1)::word))::64 Word.word) <=s ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)" unfolding repL_def repe.simps by auto then have "(0x7FFFFFFF::64 Word.word) <=s ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)" by auto then obtain r' where geq:"(r' \ r1 + r2)" and leq:"(0x7FFFFFFF \ r')" using bigThree by auto show "?thesis" unfolding repL_def repe.simps using leq geq by auto qed have int_case:"\(((scast POS_INF)::64 Word.word) <=s ?sum) \ \ (?sum <=s ((scast NEG_INF)::64 Word.word)) \ ((scast ?sum)::word) \\<^sub>L r1 + r2" proof - assume bound1:"\ ((scast POS_INF)::64 Word.word) <=s scast w1 + scast w2" assume bound2:"\ scast w1 + scast w2 <=s ((scast NEG_INF)::64 Word.word)" obtain r'::real where rDef:"r' = (real_of_int (sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)))" and r12:"r'\r1 + r2" and boundU:"r' < 0x7FFFFFFF" and boundL:" (-0x7FFFFFFF) < r'" using bigTwo[OF bound1 bound2] by auto obtain w::word where wdef:"w = (scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word))::word)" by auto then have wr:"r' = (real_of_int (sint w))" using r12 bound1 bound2 rDef by blast show ?thesis unfolding repL_def repe.simps using r12 wdef rDef boundU boundL wr by auto qed have "(let sum = ?sum in if scast POS_INF <=s sum then POS_INF else if sum <=s scast NEG_INF then NEG_INF else scast sum) \\<^sub>L r1 + r2" apply(cases "((scast POS_INF)::64 Word.word) <=s ?sum") apply(cases "?sum <=s scast NEG_INF") using inf_case neg_inf_case int_case by (auto simp add: Let_def) then show ?thesis using notinf1 notinf2 notneginf1 notneginf2 by(auto) qed qed subsection \Max function\ text\Maximum of w1 + w2 in 2s-complement\ fun wmax :: "word \ word \ word" where "wmax w1 w2 = (if w1 Correctness of wmax\ lemma wmax_lemma: assumes eq1:"w1 \\<^sub>E (r1::real)" assumes eq2:"w2 \\<^sub>E (r2::real)" shows "wmax w1 w2 \\<^sub>E (max r1 r2)" proof(cases rule: case_inf2[where ?w1.0=w1, where ?w2.0=w2]) case PosPos from PosPos eq1 eq2 have bound1:"(real_of_int (sint POS_INF)) \ r1" and bound2:"(real_of_int (sint POS_INF)) \ r2" by (auto simp add: repe.simps) have eqInf:"wmax w1 w2 = POS_INF" using PosPos unfolding wmax.simps by auto have pos_eq:"POS_INF \\<^sub>E max r1 r2" apply(rule repPOS_INF) using bound1 bound2 by linarith show ?thesis using pos_eq eqInf by auto next case PosNeg from PosNeg have bound1:"(real_of_int (sint POS_INF)) \ r1" and bound2:"r2 \ (real_of_int (sint NEG_INF))" using eq1 eq2 by (auto simp add: repe.simps) have eqNeg:"wmax w1 w2 = POS_INF" unfolding eq1 eq2 wmax.simps PosNeg word_sless_def word_sle_def by(auto) have neg_eq:"POS_INF \\<^sub>E max r1 r2" apply(rule repPOS_INF) using bound1 bound2 eq1 eq2 by auto show "?thesis" using eqNeg neg_eq by auto next case PosNum from PosNum eq1 eq2 have bound1:" (real_of_int (sint POS_INF)) \ r1" and bound2a:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" and bound2b:" (real_of_int (sint w2)) < (real_of_int (sint POS_INF))" by (auto simp add: repe.simps) have eqNeg:"wmax w1 w2 = POS_INF" using PosNum bound2b unfolding wmax.simps word_sless_def word_sle_def by auto have neg_eq:"POS_INF \\<^sub>E max r1 r2" apply (rule repPOS_INF) using bound1 bound2a bound2b word_sless_alt le_max_iff_disj unfolding eq1 eq2 by blast show "?thesis" using eqNeg neg_eq by auto next case NegPos from NegPos eq1 eq2 have bound1:"r1 \ (real_of_int (sint NEG_INF))" and bound2:" (real_of_int (sint POS_INF)) \ r2" by (auto simp add: repe.simps) have eqNeg:"wmax w1 w2 = POS_INF" unfolding NegPos word_sless_def word_sle_def by(auto) have neg_eq:"POS_INF \\<^sub>E max r1 r2" apply(rule repPOS_INF) using bound1 bound2 by auto show "wmax w1 w2 \\<^sub>E max r1 r2" using eqNeg neg_eq by auto next case NegNeg from NegNeg eq1 eq2 have bound1:"r1 \ (real_of_int (sint NEG_INF))" and bound2:"r2 \ (real_of_int (sint NEG_INF))" by (auto simp add: repe.simps) have eqNeg:"NEG_INF \\<^sub>E max r1 r2" apply(rule repNEG_INF) using eq1 eq2 bound1 bound2 by(auto) have neg_eq:"wmax w1 w2 = NEG_INF" using NegNeg by auto show "wmax w1 w2 \\<^sub>E max r1 r2" using eqNeg neg_eq by auto next case NegNum from NegNum eq1 eq2 have eq3:"r2 = (real_of_int (sint w2))" and bound2a:"(real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and bound2b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" and bound1:"r1 \ (real_of_int (sint NEG_INF))" by (auto simp add: repe.simps) have eqNeg:"max r1 r2 = (real_of_int (sint w2))" using NegNum assms(2) bound2a eq3 repeInt_simps bound1 bound2a bound2b by (metis less_irrefl max.bounded_iff max_def not_less) then have extra_eq:"(wmax w1 w2) = w2" using assms(2) bound2a eq3 NegNum repeInt_simps by (simp add: word_sless_alt) have neg_eq:"wmax w1 w2 \\<^sub>E (real_of_int (sint (wmax w1 w2)))" apply(rule repINT) using extra_eq eq3 bound2a bound2b by(auto) show "wmax w1 w2 \\<^sub>E max r1 r2" using eqNeg neg_eq extra_eq by auto next case NumPos from NumPos eq1 eq2 have p2:"w2 = POS_INF" and eq1:"r1 = (real_of_int (sint w1))" and eq2:"r2 = r2" and bound1a:"(real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2:"(real_of_int (sint POS_INF)) \ r2" by (auto simp add: repe.simps) have res1:"wmax w1 w2 = POS_INF" using NumPos p2 eq1 eq2 assms(1) bound1b p2 repeInt_simps by (simp add: word_sless_alt) have res3:"POS_INF \\<^sub>E max r1 r2" using repPOS_INF NumPos bound2 le_max_iff_disj by blast show "wmax w1 w2 \\<^sub>E max r1 r2" using res1 res3 by auto next case NumNeg from NumNeg eq1 eq2 have n2:"w2 = NEG_INF" and rw1:"r1 = (real_of_int (sint w1))" and bound1a:"(real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2:"r2 \ (real_of_int (sint NEG_INF))" by (auto simp add: repe.simps) have res1:"max r1 r2 = (real_of_int (sint (wmax w1 w2)))" using bound1b bound2 NumNeg less_trans wmax.simps of_int_less_iff word_sless_alt rw1 antisym_conv2 less_imp_le max_def by metis have res2:"wmax w1 w2 \\<^sub>E (real_of_int (sint (wmax w1 w2)))" apply(rule repINT) using bound1a bound1b bound2 NumNeg leD leI less_trans n2 wmax.simps by (auto simp add: word_sless_alt) show "wmax w1 w2 \\<^sub>E max r1 r2" using res1 res2 by auto next case NumNum from NumNum eq1 eq2 have eq1:"r1 = (real_of_int (sint w1))" and eq2:"r2 = (real_of_int (sint w2))" and bound1a:"(real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2a:"(real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and bound2b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" by (auto simp add: repe.simps) have res1:"max r1 r2 = (real_of_int (sint (wmax w1 w2)))" using eq1 eq2 bound1a bound1b bound2a bound2b by (simp add: max_def word_sless_alt) have res2:"wmax w1 w2 \\<^sub>E (real_of_int (sint (wmax w1 w2)))" apply (rule repINT) using bound1a bound1b bound2a bound2b by (simp add: \max r1 r2 = (real_of_int (sint (wmax w1 w2)))\ eq2 min_less_iff_disj)+ show "wmax w1 w2 \\<^sub>E max r1 r2" using res1 res2 by auto qed lemma max_repU1: assumes "w1 \\<^sub>U x" assumes "w2 \\<^sub>U y" shows "wmax w1 w2 \\<^sub>U x " using wmax_lemma assms repU_def by (meson le_max_iff_disj) lemma max_repU2: assumes "w1 \\<^sub>U y" assumes "w2 \\<^sub>U x" shows "wmax w1 w2 \\<^sub>U x" using wmax_lemma assms repU_def by (meson le_max_iff_disj) text\Product of w1 * w2 with bounds checking\ fun wtimes :: "word \ word \ word" where "wtimes w1 w2 = (if w1 = POS_INF \ w2 = POS_INF then POS_INF else if w1 = NEG_INF \ w2 = POS_INF then NEG_INF else if w1 = POS_INF \ w2 = NEG_INF then NEG_INF else if w1 = NEG_INF \ w2 = NEG_INF then POS_INF else if w1 = POS_INF \ w2 0 0 = w2 then 0 else if w1 = NEG_INF \ w2 0 0 = w2 then 0 else if w1 w2 = POS_INF then NEG_INF else if 0 w2 = POS_INF then POS_INF else if 0 = w1 \ w2 = POS_INF then 0 else if w1 w2 = NEG_INF then POS_INF else if 0 w2 = NEG_INF then NEG_INF else if 0 = w1 \ w2 = NEG_INF then 0 else (let prod::64 Word.word = (scast w1) * (scast w2) in if prod <=s (scast NEG_INF) then NEG_INF else if (scast POS_INF) <=s prod then POS_INF else (scast prod)))" subsection \Multiplication upper bound\ text\Product of 32-bit numbers fits in 64 bits\ lemma times_upcast_lower: fixes x y::int assumes a1:"x \ -2147483648" assumes a2:"y \ -2147483648" assumes a3:"x \ 2147483648" assumes a4:"y \ 2147483648" shows "- 4611686018427387904 \ x * y" proof - let ?thesis = "- 4611686018427387904 \ x * y" have is_neg:"- 4611686018427387904 < (0::int)" by auto have case1:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"x * y \ 0" using a5 a6 by (simp add: zero_le_mult_iff) then show ?thesis using is_neg by auto qed have case2:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"-2147483648 * (2147483648) \ x * 2147483648" using a1 a2 a3 a4 a5 a6 by linarith have h2:"-2147483648 \ y" using a6 by auto have h3:"x * 2147483648 \ x * y" using a1 a2 a3 a4 a5 a6 h2 using mult_left_mono_neg by blast show ?thesis using h1 h2 h3 by auto qed have case3:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"2147483648 * (-2147483648) \ 2147483648 * y" using a1 a2 a3 a4 a5 a6 by linarith have h2:"-2147483648 \ x" using a5 by auto have h3:"2147483648 * y \ x * y" using a1 a2 a3 a4 a5 a6 h2 using mult_left_mono_neg mult_right_mono_neg by blast show ?thesis using h1 h2 h3 by auto qed have case4:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"x * y \ 0" using a5 a6 by (simp add: zero_le_mult_iff) then show ?thesis using is_neg by auto qed show ?thesis using case1 case2 case3 case4 by linarith qed text\Product of 32-bit numbers fits in 64 bits\ lemma times_upcast_upper: fixes x y ::int assumes a1:"x \ -2147483648" assumes a2:"y \ -2147483648" assumes a3:"x \ 2147483648" assumes a4:"y \ 2147483648" shows "x * y \ 4611686018427387904" proof - let ?thesis = "x * y \ 4611686018427387904" have case1:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"2147483648 * 2147483648 \ x * 2147483648" using a1 a2 a3 a4 a5 a6 by linarith have h2:"x * 2147483648 \ x * y" using a1 a2 a3 a4 a5 a6 by (simp add: mult_mono) show ?thesis using h1 h2 by auto qed have case2:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"2147483648 * 2147483648 \ (0::int)" by auto have h2:"0 \ x * y" using a5 a6 mult_nonneg_nonpos2 by blast show ?thesis using h1 h2 by auto qed have case3:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"2147483648 * 2147483648 \ (0::int)" by auto have h2:"0 \ x * y" using a5 a6 mult_nonneg_nonpos by blast show ?thesis using h1 h2 by auto qed have case4:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"-2147483648 * -2147483648 \ x * -2147483648" using a1 a2 a3 a4 a5 a6 by linarith have h2:"x * -2147483648 \ x * y" using a1 a2 a3 a4 a5 a6 mult_left_mono_neg by blast show ?thesis using h1 h2 by auto qed show "x * y \ 4611686018427387904" using case1 case2 case3 case4 by linarith qed text\Correctness of 32x32 bit multiplication\ subsection \Exact multiplication\ lemma wtimes_exact: assumes eq1:"w1 \\<^sub>E r1" assumes eq2:"w2 \\<^sub>E r2" shows "wtimes w1 w2 \\<^sub>E r1 * r2" proof - have POS_cast:"sint ((scast POS_INF)::64 Word.word) = sint POS_INF" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have POS_sint:"sint POS_INF = (2^31)-1" by auto have w1_cast:"sint ((scast w1)::64 Word.word) = sint w1" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have w2_cast:"sint ((scast w2)::64 Word.word) = sint w2" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have NEG_cast:"sint ((scast NEG_INF)::64 Word.word) = sint NEG_INF" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have rangew1:"sint ((scast w1)::64 Word.word) \ {- (2 ^ 31).. (2^31)} " using Word.word_sint.Rep[of "(w1)::32 Word.word"] sints32 len32 mem_Collect_eq POS_cast w1_cast by auto have rangew2:"sint ((scast w2)::64 Word.word) \ {- (2 ^ 31).. (2^31)} " using Word.word_sint.Rep[of "(w2)::32 Word.word"] sints32 len32 mem_Collect_eq POS_cast w2_cast by auto show ?thesis proof (cases rule: case_times_inf[of w1 w2]) case PosPos then have a1: "PosInf \ r1" and a2: "PosInf \ r2" using "PosPos" eq1 eq2 repe.simps by (auto) have f3: "\n e::real. 1 \ max ( (numeral n)) e" by (simp add: le_max_iff_disj) have "\n e::real. 0 \ max ( (numeral n)) e" by (simp add: le_max_iff_disj) then have "r1 \ r1 * r2" using f3 "PosPos" eq1 eq2 repe.simps using eq1 eq2 by (auto simp add: repe.simps) then have "PosInf \ r1 * r2" using a1 by linarith then show ?thesis using "PosPos" by(auto simp add: repe.simps) next case NegPos from "NegPos" have notPos:"w1 \ POS_INF" unfolding POS_INF_def NEG_INF_def by auto have a1: "NegInf \ r1" using eq1 "NegPos" by (auto simp add: repe.simps) have a2: "PosInf \ r2" using eq2 "NegPos" by (auto simp add: repe.simps) have f1: "real_of_int Numeral1 = 1" by simp have f3: "(real_of_int 3) \ - r1" using a1 by (auto) have f4:"0 \ r2" using f1 a2 by(auto) have f5: "r1 \ - 1" using f3 by auto have fact:"r1 * r2 \ - r2" using f5 f4 mult_right_mono by fastforce show ?thesis using a1 a2 fact by (auto simp add: repe.simps "NegPos") next case PosNeg have a1: "PosInf \ r1" using eq1 "PosNeg" by (auto simp add: repe.simps) then have h1:"r1 \ 1" by (auto) have a2: " NegInf \ r2" using eq2 "PosNeg" by (auto simp add: repe.simps) have f1: "\ NegInf * (- 1) \ 1" by (auto) have f2: "\e ea::real. (e * (- 1) \ ea) = (ea * (- 1) \ e)" by force then have f3: "\ 1 * (- 1::real) \ NegInf" using f1 by blast have f4: "r1 * (- 1) \ NegInf" using f2 a1 by(auto) have f5: "\e ea eb. (if (ea::real) \ eb then e \ eb else e \ ea) = (e \ ea \ e \ eb)" by force have " 0 * (- 1::real) \ 1" by simp then have "r1 * (- 1) \ 0" using f5 f4 f3 f2 by meson then have f6: "0 \ r1" by fastforce have "1 * (- 1) \ (- 1::real)" using f2 by force then have fact:"r2 \ (- 1)" using f3 a2 by fastforce have rule:"\c. c > 0 \ r1 \ c \ r2 \ -1 \ r1 * r2 \ -c" apply auto by (metis (no_types, hide_lams) f5 mult_less_cancel_left_pos mult_minus1_right neg_le_iff_le not_less) have "r1 * r2 \ NegInf" using "PosNeg" f6 fact rule[of PosInf] a1 by(auto) then show ?thesis using "PosNeg" by (auto simp add: repe.simps) next case NegNeg have a1: "(-2147483647) \ r1" using eq1 "NegNeg" by (auto simp add: repe.simps) then have h1:"r1 \ -1" using max.bounded_iff max_def one_le_numeral by auto have a2: " (-2147483647) \ r2" using eq2 "NegNeg" by (auto simp add: repe.simps) have f1: "\e ea eb. \ (e::real) \ ea \ \ 0 \ eb \ eb * e \ eb * ea" using mult_left_mono by metis have f2: "- 1 = (- 1::real)" by force have f3: " 0 \ (1::real)" by simp have f4: "\e ea eb. (ea::real) \ e \ \ ea \ eb \ \ eb \ e" by (meson less_le_trans not_le) have f5: " 0 \ (2147483647::real)" by simp have f6: "- (- 2147483647) = (2147483647::real)" by force then have f7: "- ( (- 2147483647) * r1) = (2147483647 * r1)" by (metis mult_minus_left) have f8: "- ( (- 2147483647) * (- 1)) = 2147483647 * (- 1::real)" by simp have " 2147483647 = - 1 * (- 2147483647::real)" by simp then have f9: "r1 \ (- 1) \ 2147483647 \ r1 * (- 2147483647)" using f8 f7 f5 f2 f1 by linarith have f10: "- 2147483647 = (- 2147483647::real)" by fastforce have f11: "- (r2 * 1 * (r1 * (- 1))) = r1 * r2" by (simp add: mult.commute) have f12: "r1 * (- 1) = - (r1 * 1)" by simp have "r1 * 1 * ( (- 2147483647) * 1) = (- 2147483647) * r1" by (simp add: mult.commute) then have f13: "r1 * (- 1) * ( (- 2147483647) * 1) = 2147483647 * r1" using f12 f6 by (metis (no_types) mult_minus_left) have " 1 * r1 \ 1 * (- 2147483647)" using a1 by (auto simp add: a1) then have " 2147483647 \ r1 * (- 1)" by fastforce then have "0 \ r1 * (- 1)" using f5 f4 by (metis) then have "r1 \ (- 1) \ - (r1 * 2147483647) \ - (r2 * 1 * (r1 * (- 1)))" by (metis a2 f11 h1 mult_left_mono_neg minus_mult_right mult_minus1_right neg_0_le_iff_le) then have "r1 \ (- 1) \ r1 * (- 2147483647) \ r2 * r1" using f11 f10 by (metis mult_minus_left mult.commute) then have fact:" 2147483647 \ r2 * r1" using f9 f4 by blast show ?thesis using a1 a2 h1 fact by (auto simp add: repe.simps "NegNeg" mult.commute) next case PosLo from "PosLo" have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" by (auto) from eq1 "PosLo" have upper:" (real_of_int (sint POS_INF)) \ r1 " by (auto simp add: repe.simps) have lower1:"sint w2 < 0" using "PosLo" apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) then have lower2:"sint w2 \ -1" by auto from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps "PosLo" by (auto simp add: repe.simps) have f4: "r1 * (- 1) \ (- 2147483647)" using upper by (auto) have f5: "r2 \ (- 1)" using lower2 rw2 by simp have "0 < r1" using upper by (auto) have "\r. r < - 2147483647 \ \ r < r1 * - 1" using f4 less_le_trans by blast then have "r1 * (real_of_int (sint w2)) \ (- 2147483647)" using f5 f4 upper lower2 rw2 mult_left_mono by (metis \0 < r1\ dual_order.order_iff_strict f5 mult_left_mono rw2) then have "r1 * r2 \ real_of_int (sint NEG_INF)" using upper lower2 rw2 by (auto) then show ?thesis using "PosLo" by (auto simp add: repe.simps) next case PosHi from "PosHi" have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" by (auto) from eq1 "PosHi" have upper:"(real_of_int (sint POS_INF)) \ r1 " by (auto simp add: repe.simps) have lower1:"sint w2 > 0" using "PosHi" apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) then have lower2:"sint w2 \ 1" by auto from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps "PosHi" by (auto) have "0 \ r1" using upper by (auto) then have "r1 \ r1 * r2" using rw2 lower2 by (metis (no_types) mult_left_mono mult.right_neutral of_int_1_le_iff) have "PosInf \ r1 * r2" using upper lower2 rw2 apply (auto) using \0 \ r1\ mult_mono mult_numeral_1_right numeral_One of_int_1_le_iff zero_le_one by metis then show ?thesis using "PosHi" by (auto simp add: repe.simps) next case PosZero from "PosZero" have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" by (auto) from eq1 "PosZero" have upper:" (real_of_int (sint POS_INF)) \ r1 " by (auto simp add: repe.simps) have lower1:"sint w2 = 0" using "PosZero" by (auto simp add: word_sless_def word_sle_def) from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps "PosZero" by auto have "0 = r1 * r2" using "PosZero" rw2 by auto then show ?thesis using "PosZero" by (auto simp add: repe.simps) next case NegHi have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" using "NegHi" by (auto) from eq1 "NegHi" have upper:"(real_of_int (sint NEG_INF)) \ r1 " by (auto simp add: repe.simps) have low:"sint w2 > 0" using "NegHi" apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) then have lower1:"(real_of_int (sint w2)) > 0" by auto then have lower2:"(real_of_int (sint w2)) \ 1" using low by (simp add: int_le_real_less) from eq1 have rw1:"r1 \ (real_of_int (sint w1))" using repe.simps "NegHi" by (simp add: upper) from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps "NegHi" by (auto) have mylem:"\x y z::real. x \ -1 \ y \ 1 \ z \ -1 \ x \ z \ x * y \ z" proof - fix x y z::real assume a1:"x \ -1" assume a2:"y \ 1" then have h1:"-1 \ -y" by auto assume a3:"z \ -1" then have a4:"z < 0" by auto from a4 have h2:"-z > 0" using leD leI by auto from a3 have h5:"-z \ 1" by (simp add: leD leI) assume a5:"x \ z" then have h6:"-x \ -z" by auto have h3:"-x * -z = x * z" by auto show "x * y \ z" using a1 a2 a3 a5 a4 h1 h2 h3 h6 h5 a5 dual_order.trans leD mult.right_neutral by (metis dual_order.order_iff_strict mult_less_cancel_left2) qed have prereqs:"r1 \ - 1" "1 \ (real_of_int (sint w2))" " (- 2147483647::real) \ - 1 " "r1 \ (-2147483647)" using rw1 rw2 "NegHi" lower2 by (auto simp add: word_sless_def word_sle_def) have "r1 * r2 \ real_of_int (sint NEG_INF)" using upper lower1 lower2 rw1 rw2 apply (auto simp add: word_sless_def word_sle_def) using mylem[of "r1" " (real_of_int (sint w2))" " (- 2147483647)"] prereqs by auto then show ?thesis using "NegHi" by (auto simp add: repe.simps) next case NegLo from "NegLo" have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" by (auto) from eq1 "NegLo" have upper:"(real_of_int (sint NEG_INF)) \ r1" by (auto simp add: repe.simps) have low:"sint w2 < 0" using "NegLo" by (auto simp add: word_sless_def word_sle_def dual_order.order_iff_strict) then have lower1:"(real_of_int (sint w2)) < 0" by auto then have lower2:"(real_of_int (sint w2)) \ -1" using low by (simp add: int_le_real_less) from eq1 have rw1:"r1 \ (real_of_int (sint w1))" using repe.simps "NegLo" by (simp add: upper) from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps "NegLo" by (auto) have hom:"(- 2147483647) = -(2147483647::real)" by auto have mylem:"\x y z::real. y < 0 \ x \ y \ z \ -1 \ -y \ x * z" proof - fix x y z::real assume a1:"y < 0" assume a2:"x \ y" then have h1:"-x \ -y" by auto assume a3:"z \ -1" then have a4:"z < 0" by auto from a4 have h2:"-z > 0" using leD leI by auto from a3 have h5:"-z \ 1" by (simp add: leD leI) have h4:"-x * -z \ -y" using a1 a2 a3 a4 h1 h2 h5 dual_order.trans mult.right_neutral by (metis mult.commute neg_0_less_iff_less real_mult_le_cancel_iff1) have h3:"-x * -z = x * z" by auto show "- y \ x * z " using a1 a2 a3 a4 h1 h2 h3 h4 h5 by simp qed have prereqs:"- 2147483647 < (0::real)" " r1 \ - 2147483647" " (real_of_int (sint w2)) \ - 1" using rw1 rw2 "NegLo" by (auto simp add: word_sless_def word_sle_def lower2) have "2147483647 \ r1 * r2" using upper lower1 lower2 rw1 rw2 prereqs mylem[of "-2147483647" "r1" "(real_of_int (sint w2))"] by (auto simp add: word_sless_def word_sle_def) then show ?thesis using "NegLo" by (auto simp add: repe.simps) next case NegZero from "NegZero" have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" by (auto) from eq2 "NegZero" have "r2 = 0" using repe.simps "NegZero" by (auto) then show ?thesis using "NegZero" by (auto simp add: repe.simps) next case LoPos from "LoPos" have w2NotPinf:"w1 \ POS_INF" and w2NotNinf:"w1 \ NEG_INF" by (auto) from eq2 "LoPos" have upper:"(real_of_int (sint POS_INF)) \ r2 " by (auto simp add: repe.simps) have lower1:"sint w1 < 0" using "LoPos" apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) then have lower2:"sint w1 \ -1" by auto from eq1 have rw1:"r1 = (real_of_int (sint w1))" using repe.simps "LoPos" by (auto simp add: repe.simps) have f4: "r2 * (- 1) \ (- 2147483647)" using upper by(auto) have f5: "r1 \ (- 1)" using lower2 rw1 by simp have "0 < r2" using upper by(auto) then have "r2 * r1 \ r2 * (- 1)" by (metis dual_order.order_iff_strict mult_right_mono f5 mult.commute) then have "r2 * r1 \ (- 2147483647)" by (meson f4 less_le_trans not_le) then have "(real_of_int (sint w1)) * r2 \ (- 2147483647)" using f5 f4 rw1 less_le_trans not_le mult.commute rw1 by (auto simp add: mult.commute) then have "r1 * r2 \ NegInf" using rw1 by (auto) then show ?thesis using "LoPos" by (auto simp: repe.simps) next case HiPos from "HiPos" have w2NotPinf:"w1 \ POS_INF" and w2NotNinf:"w1 \ NEG_INF" by (auto) from eq2 "HiPos" have upper:"(real_of_int (sint POS_INF)) \ r2 " by (auto simp add: repe.simps) have lower1:"sint w1 > 0" using "HiPos" by (auto simp add: word_sless_def word_sle_def dual_order.order_iff_strict) then have lower2:"sint w1 \ 1" by auto from eq1 have rw2:"r1 = (real_of_int (sint w1))" using "HiPos" by (auto simp add: repe.simps) have "0 \ r2" using upper by(auto) then have "r2 \ r2 * r1" using lower2 rw2 by (metis (no_types) mult_left_mono mult.right_neutral of_int_1_le_iff) have "2147483647 \ r1 * r2" using upper lower2 rw2 by (auto simp add: word_sless_def word_sle_def order_trans) then show ?thesis using "HiPos" by (auto simp add: repe.simps) next case ZeroPos from "ZeroPos" have w2NotPinf:"w1 \ POS_INF" and w2NotNinf:"w1 \ NEG_INF" by (auto) from eq2 "ZeroPos" have upper:" (real_of_int (sint POS_INF)) \ r2 " by (auto simp add: repe.simps) have lower1:"sint w1 = 0" using "ZeroPos" by (auto simp add: word_sless_def word_sle_def) from eq1 have rw2:"r1 = (real_of_int (sint w1))" using repe.simps "ZeroPos" by (auto) have "r1 = 0" using lower1 rw2 by auto then show ?thesis using "ZeroPos" by (auto simp add: repe.simps) next case ZeroNeg from "ZeroNeg" have w2NotPinf:"w1 \ POS_INF" and w2NotNinf:"w1 \ NEG_INF" by (auto) from eq2 "ZeroNeg" have upper:"(real_of_int (sint NEG_INF)) \ r2 " by (auto simp add: repe.simps) have lower1:"sint w1 = 0" using "ZeroNeg" by (auto simp add: word_sless_def word_sle_def) from eq1 have rw2:"r1 = (real_of_int (sint w1))" using repe.simps "ZeroNeg" by (auto) have "r1 = 0" using lower1 rw2 by auto then show ?thesis using "ZeroNeg" by (auto simp add: repe.simps) next case LoNeg from "LoNeg" have w2NotPinf:"w1 \ POS_INF" and w2NotNinf:"w1 \ NEG_INF" by (auto) from eq2 "LoNeg" have upper:" (real_of_int (sint NEG_INF)) \ r2 " by (auto simp add: repe.simps) have low:"sint w1 < 0" using "LoNeg" apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) then have lower1:"(real_of_int (sint w1)) < 0" by auto then have lower2:"(real_of_int (sint w1)) \ -1" using low by (simp add: int_le_real_less) from eq1 have rw1:"r2 \ (real_of_int (sint w2))" using "LoNeg" upper by auto from eq1 have rw2:"r1 = (real_of_int (sint w1))" using "LoNeg" by (auto simp add: upper repe.simps) have hom:"(- 2147483647::real) = -(2147483647)" by auto have mylem:"\x y z::real. y < 0 \ x \ y \ z \ -1 \ -y \ x * z" proof - fix x y z::real assume a1:"y < 0" assume a2:"x \ y" then have h1:"-x \ -y" by auto assume a3:"z \ -1" then have a4:"z < 0" by auto from a4 have h2:"-z > 0" using leD leI by auto from a3 have h5:"-z \ 1" by (simp add: leD leI) have h4:"-x * -z \ -y" using a1 a2 a3 a4 h1 h2 h5 dual_order.trans mult_left_mono mult.right_neutral mult.commute by (metis dual_order.order_iff_strict mult_minus_right mult_zero_right neg_le_iff_le) have h3:"-x * -z = x * z" by auto show "- y \ x * z " using a1 a2 a3 a4 h1 h2 h3 h4 h5 by simp qed have prereqs:"- 2147483647 < (0::real)" " r2 \ - 2147483647" " (real_of_int (sint w1)) \ - 1" using rw1 rw2 "LoNeg" lower2 by (auto simp add: word_sless_def word_sle_def lower2) have "2147483647 \ r1 * r2" using upper lower1 lower2 rw1 rw2 mylem[of "-2147483647" "r2" "(real_of_int (sint w1))"] prereqs by (auto simp add:word_sless_def word_sle_def mult.commute) then show ?thesis using "LoNeg" by (auto simp add: repe.simps) next case HiNeg from HiNeg have w1NotPinf:"w1 \ POS_INF" and w1NotNinf:"w1 \ NEG_INF" by (auto) have upper:" (real_of_int (sint NEG_INF)) \ r2 " using HiNeg eq2 by (auto simp add: repe.simps ) have low:"sint w1 > 0" using HiNeg apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) then have lower1:"(real_of_int (sint w1)) > 0" by auto then have lower2:"(real_of_int (sint w1)) \ 1" using low by (simp add: int_le_real_less) from eq2 have rw1:"r2 \ (real_of_int (sint w2))" using repe.simps HiNeg by (simp add: upper) from eq1 have rw2:"r1 = (real_of_int (sint w1))" using repe.simps HiNeg by (auto) have mylem:"\x y z::real. x \ -1 \ y \ 1 \ z \ -1 \ x \ z \ x * y \ z" proof - fix x y z::real assume a1:"x \ -1" assume a2:"y \ 1" then have h1:"-1 \ -y" by auto assume a3:"z \ -1" then have a4:"z < 0" by auto from a4 have h2:"-z > 0" using leD leI by auto from a3 have h5:"-z \ 1" by (simp add: leD leI) assume a5:"x \ z" then have h6:"-x \ -z" by auto have h3:"-x * -z = x * z" by auto show "x * y \ z" using a1 a2 a3 a4 h1 h2 h3 h6 h5 a5 dual_order.trans less_eq_real_def by (metis mult_less_cancel_left1 not_le) qed have prereqs:"r2 \ - 1" "1 \ (real_of_int (sint w1))" " (- 2147483647) \ - (1::real )" "r2 \ (- 2147483647)" using rw1 rw2 HiNeg lower2 by (auto simp add: word_sless_def word_sle_def) have "r1 * r2 \ - 2147483647" using upper lower1 lower2 rw1 rw2 apply (auto simp add: word_sless_def word_sle_def) using mylem[of "r2" "(real_of_int (sint w1))" " (- 2147483647)"] prereqs by (auto simp add: mult.commute) then show ?thesis using HiNeg by(auto simp add: repe.simps) next case AllFinite let ?prod = "(((scast w1)::64 Word.word) * ((scast w2)::64 Word.word))" consider (ProdNeg) "?prod <=s ((scast NEG_INF)::64 Word.word)" | (ProdPos) "(((scast POS_INF)::64 Word.word) <=s ?prod)" | (ProdFin) "\(?prod <=s ((scast NEG_INF)::64 Word.word)) \ \((scast POS_INF)::64 Word.word) <=s ?prod" by (auto) then show ?thesis proof (cases) case ProdNeg have bigLeq:"(4611686018427387904::real) \ 9223372036854775807" by auto have set_cast:"\x::int. (x \ {-(2^31)..2^31}) = ( (real_of_int x) \ {-(2^31)..2^31})" by auto have eq3:"sint(((scast w1)::64 Word.word) * ((scast w2)::64 Word.word)) = sint ((scast w1)::64 Word.word) * sint ((scast w2)::64 Word.word)" apply(rule Word_Lemmas.signed_arith_sint(4)) using rangew1 rangew2 w1_cast w2_cast using Word.word_size[of "((scast w1)::64 Word.word)"] using Word.word_size[of "((scast w2)::64 Word.word)"] using times_upcast_upper[of "sint w1" "sint w2"] using times_upcast_lower[of "sint w1" "sint w2"] by auto assume "?prod <=s ((scast NEG_INF)::64 Word.word)" then have sint_leq:"sint ?prod \ sint ((scast NEG_INF)::64 Word.word)" using word_sle_def by blast have neqs:"w1 \ POS_INF" " w1 \ NEG_INF" "w2 \ POS_INF" "w2 \ NEG_INF" using AllFinite word_sless_def signed.not_less_iff_gr_or_eq by force+ from eq1 have rw1:"r1 = (real_of_int (sint w1))" using neqs by (auto simp add: repe.simps) from eq2 have rw2:"r2 = (real_of_int (sint w2))" using neqs by (auto simp add: repe.simps) show ?thesis using AllFinite ProdNeg w1_cast w2_cast rw1 rw2 sint_leq apply (auto simp add: repe.simps) by (metis (no_types, hide_lams) eq3 of_int_le_iff of_int_minus of_int_mult of_int_numeral) next case ProdPos have bigLeq:"(4611686018427387904::real) \ 9223372036854775807" by auto have set_cast:"\x::int. (x \ {-(2^31)..2^31}) = ( (real_of_int x) \ {-(2^31)..2^31})" by auto have eq3:"sint(((scast w1)::64 Word.word) * ((scast w2)::64 Word.word)) = sint ((scast w1)::64 Word.word) * sint ((scast w2)::64 Word.word)" apply(rule Word_Lemmas.signed_arith_sint(4)) using rangew1 rangew2 POS_cast POS_sint w1_cast w2_cast using Word.word_size[of "((scast w1)::64 Word.word)"] using Word.word_size[of "((scast w2)::64 Word.word)"] using times_upcast_upper[of "sint w1" "sint w2"] using times_upcast_lower[of "sint w1" "sint w2"] by auto assume cast:"((scast POS_INF)::64 Word.word) <=s ?prod" then have sint_leq:"sint ((scast POS_INF)::64 Word.word) \ sint ?prod" using word_sle_def by blast have neqs:"w1 \ POS_INF" " w1 \ NEG_INF" "w2 \ POS_INF" "w2 \ NEG_INF" using AllFinite word_sless_def signed.not_less_iff_gr_or_eq by force+ from eq1 have rw1:"r1 = (real_of_int (sint w1))" using repe.simps AllFinite neqs by auto from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps AllFinite neqs by auto have prodHi:"r1 * r2 \ PosInf" using w1_cast w2_cast rw1 rw2 sint_leq apply(auto) by (metis (no_types, hide_lams) eq3 of_int_le_iff of_int_mult of_int_numeral) have infs:"SCAST(32 \ 64) NEG_INF 64) POS_INF" by (auto) have casted:"SCAST(32 \ 64) POS_INF <=s SCAST(32 \ 64) w1 * SCAST(32 \ 64) w2" using cast by auto have almostContra:"SCAST(32 \ 64) NEG_INF 64) w1 * SCAST(32 \ 64) w2" using infs cast signed.order.strict_trans2 by blast have contra:"\(SCAST(32 \ 64) w1 * SCAST(32 \ 64) w2 <=s SCAST(32 \ 64) NEG_INF)" using eq3 almostContra by auto have wtimesCase:"wtimes w1 w2 = POS_INF" using neqs ProdPos almostContra wtimes.simps AllFinite ProdPos by (auto simp add: repe.simps Let_def) show ?thesis using prodHi apply(simp only: repe.simps) apply(rule disjI1) apply(rule exI[where x= "r1*r2"]) apply(rule conjI) apply(rule wtimesCase) using prodHi by auto next case ProdFin have bigLeq:"(4611686018427387904::real) \ 9223372036854775807" by auto have set_cast:"\x::int. (x \ {-(2^31)..2^31}) = ( (real_of_int x) \ {-(2^31)..2^31})" by auto have eq3:"sint(((scast w1)::64 Word.word) * ((scast w2)::64 Word.word)) = sint ((scast w1)::64 Word.word) * sint ((scast w2)::64 Word.word)" apply(rule Word_Lemmas.signed_arith_sint(4)) using rangew1 rangew2 POS_cast POS_sint w1_cast w2_cast using Word.word_size[of "((scast w1)::64 Word.word)"] using Word.word_size[of "((scast w2)::64 Word.word)"] using times_upcast_upper[of "sint w1" "sint w2"] using times_upcast_lower[of "sint w1" "sint w2"] by auto from ProdFin have a1:"\(?prod <=s ((scast NEG_INF)::64 Word.word))" by auto then have sintGe:"sint (?prod) > sint (((scast NEG_INF)::64 Word.word))" using word_sle_def dual_order.order_iff_strict signed.linear by fastforce from ProdFin have a2:"\((scast POS_INF)::64 Word.word) <=s ?prod" by auto then have sintLe:"sint (((scast POS_INF)::64 Word.word)) > sint (?prod)" using word_sle_def dual_order.order_iff_strict signed.linear by fastforce have neqs:"w1 \ POS_INF" " w1 \ NEG_INF" "w2 \ POS_INF" "w2 \ NEG_INF" using AllFinite word_sless_def signed.not_less_iff_gr_or_eq by force+ from eq1 have rw1:"r1 = (real_of_int (sint w1))" using neqs by(auto simp add: repe.simps) from eq2 have rw2:"r2 = (real_of_int (sint w2))" using neqs by(auto simp add: repe.simps) from rw1 rw2 have "r1 * r2 = (real_of_int ((sint w1) * (sint w2)))" by simp have rightSize:"sint (((scast w1)::64 Word.word) * ((scast w2)::64 Word.word)) \ sints (len_of TYPE(32))" using sintLe sintGe sints32 by (simp) have downcast:"sint ((scast (((scast w1)::64 Word.word) * ((scast w2)::64 Word.word)))::word) = sint (((scast w1)::64 Word.word) * ((scast w2)::64 Word.word))" using scast_down_range[OF rightSize] by auto then have res_eq:"r1 * r2 = real_of_int(sint((scast (((scast w1)::64 Word.word)*((scast w2)::64 Word.word)))::word))" using rw1 rw2 eq3 POS_cast POS_sint w1_cast w2_cast downcast \r1 * r2 = (real_of_int (sint w1 * sint w2))\ by (auto) have res_up:"sint (scast (((scast w1)::64 Word.word) * ((scast w2)::64 Word.word))::word) < sint POS_INF" using rw1 rw2 eq3 POS_cast POS_sint w1_cast w2_cast downcast \r1 * r2 = (real_of_int (sint w1 * sint w2))\ \sint (scast w1 * scast w2) < sint (scast POS_INF)\ of_int_eq_iff res_eq by presburger have res_lo:"sint NEG_INF < sint (scast (((scast w1)::64 Word.word) * ((scast w2)::64 Word.word))::word)" using rw1 rw2 eq3 POS_cast POS_sint w1_cast w2_cast NEG_cast downcast \r1 * r2 = (real_of_int (sint w1 * sint w2))\ \sint (scast NEG_INF) < sint (scast w1 * scast w2)\ of_int_eq_iff res_eq by presburger have "scast ?prod \\<^sub>E (r1 * r2)" using res_eq res_up res_lo by (auto simp add: rep_simps) then show ?thesis using AllFinite ProdFin by(auto) qed qed qed subsection \Multiplication upper bound\ text\Upper bound of multiplication from upper and lower bounds\ fun tu :: "word \ word \ word \ word \ word" where "tu w1l w1u w2l w2u = wmax (wmax (wtimes w1l w2l) (wtimes w1u w2l)) (wmax (wtimes w1l w2u) (wtimes w1u w2u))" lemma tu_lemma: assumes u1:"u\<^sub>1 \\<^sub>U (r1::real)" assumes u2:"u\<^sub>2 \\<^sub>U (r2::real)" assumes l1:"l\<^sub>1 \\<^sub>L (r1::real)" assumes l2:"l\<^sub>2 \\<^sub>L (r2::real)" shows "tu l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>U (r1 * r2)" proof - obtain rl1 rl2 ru1 ru2 :: real where gru1:"ru1 \ r1" and gru2:"ru2 \ r2" and grl1:"rl1 \ r1" and grl2:"rl2 \ r2" and eru1:"u\<^sub>1 \\<^sub>E ru1" and eru2:"u\<^sub>2 \\<^sub>E ru2" and erl1:"l\<^sub>1 \\<^sub>E rl1" and erl2:"l\<^sub>2 \\<^sub>E rl2" using u1 u2 l1 l2 unfolding repU_def repL_def by auto have timesuu:"wtimes u\<^sub>1 u\<^sub>2 \\<^sub>E ru1 * ru2" using wtimes_exact[OF eru1 eru2] by auto have timesul:"wtimes u\<^sub>1 l\<^sub>2 \\<^sub>E ru1 * rl2" using wtimes_exact[OF eru1 erl2] by auto have timeslu:"wtimes l\<^sub>1 u\<^sub>2 \\<^sub>E rl1 * ru2" using wtimes_exact[OF erl1 eru2] by auto have timesll:"wtimes l\<^sub>1 l\<^sub>2 \\<^sub>E rl1 * rl2" using wtimes_exact[OF erl1 erl2] by auto have maxt12:"wmax (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2) \\<^sub>E max (rl1 * rl2) (ru1 * rl2)" by (rule wmax_lemma[OF timesll timesul]) have maxt34:"wmax (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2) \\<^sub>E max (rl1 * ru2) (ru1 * ru2)" by (rule wmax_lemma[OF timeslu timesuu]) have bigMax:"wmax (wmax (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2)) (wmax (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2)) \\<^sub>E max (max (rl1 * rl2) (ru1 * rl2)) (max (rl1 * ru2) (ru1 * ru2))" by (rule wmax_lemma[OF maxt12 maxt34]) obtain maxt12val :: real where maxU12:"wmax (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2) \\<^sub>U max (rl1 * rl2) (ru1 * rl2)" using maxt12 unfolding repU_def by blast obtain maxt34val :: real where maxU34:"wmax (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2) \\<^sub>U max (rl1 * ru2) (ru1 * ru2)" using maxt34 unfolding repU_def by blast obtain bigMaxU:"wmax (wmax (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2)) (wmax (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2)) \\<^sub>U max (max (rl1 * rl2) (ru1 * rl2)) (max (rl1 * ru2) (ru1 * ru2))" using bigMax unfolding repU_def by blast have ivl1:"rl1 \ ru1" using grl1 gru1 by auto have ivl2:"rl2 \ ru2" using grl2 gru2 by auto let ?thesis = "tu l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>U r1 * r2" show ?thesis using ivl1 ivl2 proof(cases rule: case_ivl_zero) case ZeroZero assume "rl1 \ 0 \ 0 \ ru1 \ rl2 \ 0 \ 0 \ ru2" then have geq1:"ru1 \ 0" and geq2:"ru2 \ 0" by auto consider "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" using le_cases by auto then show "tu l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>U r1 * r2" proof (cases) case 1 have g1:"ru1 * ru2 \ ru1 * r2" using "1" geq1 geq2 grl2 gru2 by (simp add: mult_left_mono) have g2:"ru1 * r2 \ r1 * r2" using "1" geq1 geq2 grl1 grl2 gru1 gru2 by (simp add: mult_right_mono) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up eru1 eru2 erl1 erl2 repU_def timesuu tu.simps max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] by (metis wmax.elims) next case 2 have g1:"ru1 * ru2 \ 0" using "2" geq1 geq2 grl2 gru2 by (simp) have g2:"0 \ r1 * r2" using "2" by (simp add: mult_le_0_iff) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.coboundedI1 max.commute maxt34 by (metis repU_def tu.simps) next case 3 have g1:"ru1 * ru2 \ 0" using "3" geq1 geq2 by simp have g2:"0 \ r1 * r2" using "3" by (simp add: mult_le_0_iff) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] repU_def tu.simps timesuu by (metis max.coboundedI1 max.commute maxt34) next case 4 have g1:"rl1 * rl2 \ rl1 * r2" using "4" geq1 geq2 grl1 grl2 gru1 gru2 using \rl1 \ 0 \ 0 \ ru1 \ rl2 \ 0 \ 0 \ ru2\ less_eq_real_def by (metis mult_left_mono_neg) have g2:"rl1 * r2 \ r1 * r2" using "4" geq1 geq2 grl1 grl2 gru1 gru2 \rl1 \ 0 \ 0 \ ru1 \ rl2 \ 0 \ 0 \ ru2\ by (metis mult_left_mono_neg mult.commute) from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU1 repU_def timesll tu.simps) qed next case ZeroPos assume bounds:"rl1 \ 0 \ 0 \ ru1 \ 0 \ rl2" have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast consider "r1 \ 0" | "r1 \ 0" using le_cases by (auto) then show ?thesis proof (cases) case 1 assume r1:"r1 \ 0" have g1:"ru1 * ru2 \ ru1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_left_mono by blast have g2:"ru1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_right_mono by blast from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.coboundedI1 max.commute maxt34 by (metis repU_def tu.simps) next case 2 assume r1:"r1 \ 0" have g1:"ru1 * ru2 \ 0" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_left_mono by (simp add: mult_less_0_iff less_le_trans not_less) have g2:"0 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_right_mono by (simp add: mult_le_0_iff) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.coboundedI1 max.commute maxt34 by (metis repU_def tu.simps) qed next case ZeroNeg assume bounds:"rl1 \ 0 \ 0 \ ru1 \ ru2 \ 0" have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast have case1:"r1 \ 0 \ ?thesis" proof - assume r1:"r1 \ 0" have g1:"rl1 * rl2 \ 0" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_less_0_iff less_le_trans not_less by metis have g2:"0 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_right_mono by (simp add: mult_le_0_iff) from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU2 max_repU1 repU_def timesll tu.simps) qed have case2:"r1 \ 0 \ ?thesis" proof - assume r1:"r1 \ 0" have g1:"rl1 * rl2 \ rl1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 by (metis mult_left_mono_neg) have g2:"rl1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult.commute by (metis mult_left_mono_neg) from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU1 repU_def timesll tu.simps) qed show "tu l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>U r1 * r2" using case1 case2 le_cases by blast next case PosZero assume bounds:"0 \ rl1 \ rl2 \ 0 \ 0 \ ru2" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast consider "r2 \ 0" | "r2 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 have g1:"ru1 * ru2 \ ru1 * r2" using "1" bounds grl1 grl2 gru1 gru2 using mult_left_mono using leD leI less_le_trans by metis have g2:"ru1 * r2 \ r1 * r2" using "1" bounds grl1 grl2 gru1 gru2 using mult_right_mono by blast from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.coboundedI1 max.commute maxt34 by (metis repU_def tu.simps) next case 2 have g1:"ru1 * ru2 \ 0" using r1 bounds grl2 gru2 gru1 leD leI less_le_trans by auto have g2:"0 \ r1 * r2" using r1 "2" by (simp add: mult_le_0_iff) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.coboundedI1 max.commute maxt34 by (metis repU_def tu.simps) qed next case NegZero assume bounds:"ru1 \ 0 \ rl2 \ 0 \ 0 \ ru2" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast consider "r2 \ 0" | "r2 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 have g1:"ru1 * rl2 \ 0" using r1 "1" bounds grl1 grl2 gru1 gru2 mult_less_0_iff not_less by metis have g2:"0 \ r1 * r2" using r1 "1" bounds grl1 grl2 gru1 gru2 by (simp add: mult_le_0_iff) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU1 repU_def timesul tu.simps) next case 2 have lower:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have g1:"rl1 * rl2 \ rl1 * r2" using r1 "2" bounds grl1 grl2 gru1 gru2 less_eq(1) less_le_trans not_less mult_le_cancel_left by metis have g2:"rl1 * r2 \ r1 * r2" using r1 "2" bounds grl1 grl2 gru1 gru2 mult.commute not_le lower mult_le_cancel_left by metis from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU1 repU_def timesll tu.simps) qed next case NegNeg assume bounds:"ru1 \ 0 \ ru2 \ 0" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast have lower1:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have lower2:"rl2 \ 0" using bounds dual_order.trans grl2 r2 by blast have g1:"rl1 * rl2 \ rl1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 less_eq(1) mult_le_cancel_left less_le_trans not_less by metis have g2:"rl1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult.commute not_le lower1 lower2 mult_le_cancel_left by metis from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU1 repU_def timesll tu.simps) next case NegPos assume bounds:"ru1 \ 0 \ 0 \ rl2" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast have lower1:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have lower2:"rl2 \ 0" using bounds by auto have upper1:"ru1 \ 0" using bounds by auto have upper2:"ru2 \ 0" using bounds dual_order.trans gru2 r2 by blast have g1:"ru1 * rl2 \ ru1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 not_less upper1 lower2 mult_le_cancel_left by metis have g2:"ru1 * r2 \ r1 * r2" using r1 upper1 r2 mult_right_mono gru1 by metis from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims maxt34 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] by (metis max_repU1 repU_def timesul tu.simps) next case PosNeg assume bounds:"0 \ rl1 \ ru2 \ 0" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast have lower1:"rl1 \ 0" using bounds by auto have lower2:"rl2 \ 0" using dual_order.trans grl2 r2 by blast have upper1:"ru1 \ 0" using dual_order.trans gru1 u1 r1 by blast have upper2:"ru2 \ 0" using bounds by auto have g1:"rl1 * ru2 \ rl1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 not_less upper2 lower1 mult_le_cancel_left by metis have g2:"rl1 * r2 \ r1 * r2" using r1 lower1 r2 not_less gru2 gru1 grl1 grl2 by (metis mult_le_cancel_left mult.commute) from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show "tu l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>U r1 * r2" using up maxU12 maxU34 bigMaxU wmax.elims max.coboundedI1 max.commute maxt34 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] by (metis repU_def tu.simps) next case PosPos assume bounds:"0 \ rl1 \ 0 \ rl2" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast have lower1:"rl1 \ 0" using bounds by auto have lower2:"rl2 \ 0" using bounds by auto have upper1:"ru1 \ 0" using dual_order.trans gru1 u1 r1 by blast have upper2:"ru2 \ 0" using dual_order.trans gru2 u2 r2 bounds by blast have g1:"ru1 * ru2 \ ru1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_left_mono leD leI less_le_trans by metis have g2:"ru1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_right_mono by metis from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max.coboundedI1 max.commute maxt34 max_repU2[OF bigMaxU] max_repU2[OF maxU12] max_repU2[OF maxU34] by (metis repU_def tu.simps) qed qed subsection \Minimum function\ text\Minimum of 2s-complement words\ fun wmin :: "word \ word \ word" where "wmin w1 w2 = (if w1 Correctness of wmin\ lemma wmin_lemma: assumes eq1:"w1 \\<^sub>E (r1::real)" assumes eq2:"w2 \\<^sub>E (r2::real)" shows "wmin w1 w2 \\<^sub>E (min r1 r2)" proof(cases rule: case_inf2[where ?w1.0=w1, where ?w2.0=w2]) case PosPos assume p1:"w1 = POS_INF" and p2:"w2 = POS_INF" then have bound1:"(real_of_int (sint POS_INF)) \ r1" and bound2:"(real_of_int (sint POS_INF)) \ r2" using eq1 eq2 by (auto simp add: rep_simps repe.simps) have eqInf:"wmin w1 w2 = POS_INF" using p1 p2 unfolding wmin.simps by auto have pos_eq:"POS_INF \\<^sub>E min r1 r2" apply(rule repPOS_INF) using bound1 bound2 unfolding eq1 eq2 by auto show ?thesis using pos_eq eqInf by auto next case PosNeg assume p1:"w1 = POS_INF" assume n2:"w2 = NEG_INF" obtain r ra :: real where bound1:" (real_of_int (sint POS_INF)) \ r" and bound2:"ra \ (real_of_int (sint NEG_INF))" and eq1:"r1 = r" and eq2:"r2 = ra" using p1 n2 eq1 eq2 by(auto simp add: rep_simps repe.simps) have eqNeg:"wmin w1 w2 = NEG_INF" unfolding eq1 eq2 wmin.simps p1 n2 word_sless_def word_sle_def by(auto) have neg_eq:"NEG_INF \\<^sub>E min r1 r2" apply(rule repNEG_INF) using bound1 bound2 eq1 eq2 by auto show "?thesis" using eqNeg neg_eq by auto next case PosNum assume p1:"w1 = POS_INF" assume np2:"w2 \ POS_INF" assume nn2:"w2 \ NEG_INF" have eq2:"r2 = (real_of_int (sint w2))" and bound1:"(real_of_int (sint POS_INF)) \ r1" and bound2a:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" and bound2b:"(real_of_int (sint w2)) < (real_of_int (sint POS_INF))" using p1 np2 nn2 eq1 eq2 by(auto simp add: rep_simps repe.simps) have eqNeg:"min r1 r2 = sint w2" using p1 by (metis bound1 bound2b dual_order.trans eq2 min_def not_less) have neg_eq:"wmin w1 w2 \\<^sub>E (real_of_int (sint (wmin w1 w2)))" apply (rule repINT) using bound1 bound2a bound2b bound2b p1 unfolding eq1 eq2 by (auto simp add: word_sless_alt) show "?thesis" using eqNeg neg_eq by (metis bound2b less_eq_real_def not_less of_int_less_iff p1 wmin.simps word_sless_alt) next case NegPos assume n1:"w1 = NEG_INF" assume p2:"w2 = POS_INF" have bound1:"r1 \ (real_of_int (sint NEG_INF))" and bound2:"(real_of_int (sint POS_INF)) \ r2" using n1 p2 eq1 eq2 by(auto simp add: rep_simps repe.simps) have eqNeg:"wmin w1 w2 = NEG_INF" unfolding eq1 eq2 wmin.simps n1 p2 word_sless_def word_sle_def by(auto) have neg_eq:"NEG_INF \\<^sub>E min r1 r2" apply(rule repNEG_INF) using bound1 bound2 unfolding eq1 eq2 by auto show "wmin w1 w2 \\<^sub>E min r1 r2" using eqNeg neg_eq by auto next case NegNeg assume n1:"w1 = NEG_INF" assume n2:"w2 = NEG_INF" have bound1:"r1 \ (real_of_int (sint NEG_INF))" and bound2:"r2 \ (real_of_int (sint NEG_INF))" using n1 n2 eq1 eq2 by(auto simp add: rep_simps repe.simps) have eqNeg:"NEG_INF \\<^sub>E min r1 r2" apply(rule repNEG_INF) using eq1 eq2 bound1 bound2 unfolding NEG_INF_def by (auto) have neg_eq:"wmin w1 w2 = NEG_INF" using n1 n2 unfolding NEG_INF_def wmin.simps by auto show "wmin w1 w2 \\<^sub>E min r1 r2" using eqNeg neg_eq by auto next case NegNum assume n1:"w1 = NEG_INF" and nn2:"w2 \ NEG_INF" and np2:"w2 \ POS_INF" have eq2:"r2 = (real_of_int (sint w2))" and bound2a:"(real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and bound2b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" and bound1:"r1 \ (real_of_int (sint NEG_INF))" using n1 nn2 np2 eq2 eq1 eq2 by (auto simp add: rep_simps repe.simps) have eqNeg:"wmin w1 w2 = NEG_INF" using n1 assms(2) bound2a eq2 n1 repeInt_simps by (auto simp add: word_sless_alt) have neg_eq:"NEG_INF \\<^sub>E min r1 r2" apply(rule repNEG_INF) using bound1 bound2a bound2b eq1 min_le_iff_disj by blast show "wmin w1 w2 \\<^sub>E min r1 r2" using eqNeg neg_eq by auto next case NumPos assume p2:"w2 = POS_INF" and nn1:"w1 \ NEG_INF" and np1:"w1 \ POS_INF" have eq1:"r1 = (real_of_int (sint w1))" and bound1a:" (real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2:" (real_of_int (sint POS_INF)) \ r2" using nn1 np1 p2 eq2 eq1 eq2 by (auto simp add: rep_simps repe.simps) have res1:"wmin w1 w2 = w1" using p2 eq1 eq2 assms(1) bound1b p2 repeInt_simps by (auto simp add: word_sless_alt) have res2:"min r1 r2 = (real_of_int (sint w1))" using eq1 eq2 bound1a bound1b bound2 by (auto simp add: less_imp_le less_le_trans min_def) have res3:"wmin w1 w2 \\<^sub>E (real_of_int (sint (wmin w1 w2)))" apply(rule repINT) using p2 bound1a res1 bound1a bound1b bound2 by auto show "wmin w1 w2 \\<^sub>E min r1 r2" using res1 res2 res3 by auto next case NumNeg assume nn1:"w1 \ NEG_INF" assume np1:"w1 \ POS_INF" assume n2:"w2 = NEG_INF" have eq1:"r1 = (real_of_int (sint w1))" and bound1a:" (real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2:"r2 \ (real_of_int (sint NEG_INF))" using nn1 np1 n2 eq2 eq1 eq2 by (auto simp add: rep_simps repe.simps) have res1:"wmin w1 w2 = NEG_INF" using n2 bound1b by (metis min.absorb_iff2 min_def n2 not_less of_int_less_iff wmin.simps word_sless_alt) have res2:"NEG_INF \\<^sub>E min r1 r2" apply(rule repNEG_INF) using eq1 eq2 bound1a bound1b bound2 min_le_iff_disj by blast show "wmin w1 w2 \\<^sub>E min r1 r2" using res1 res2 by auto next case NumNum assume np1:"w1 \ POS_INF" assume nn1:"w1 \ NEG_INF" assume np2:"w2 \ POS_INF" assume nn2:"w2 \ NEG_INF" have eq1:"r1 = (real_of_int (sint w1))" and eq2:"r2 = (real_of_int (sint w2))" and bound1a:" (real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2a:" (real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and bound2b:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" using nn1 np1 nn2 np2 eq2 eq1 eq2 by (auto simp add: rep_simps repe.simps) have res1:"min r1 r2 = (real_of_int (sint (wmin w1 w2)))" using eq1 eq2 bound1a bound1b bound2a bound2b by (simp add: min_def word_sless_alt) have res2:"wmin w1 w2 \\<^sub>E (real_of_int (sint (wmin w1 w2)))" apply (rule repINT) using bound1a bound1b bound2a bound2b by (simp add: \min r1 r2 = (real_of_int (sint (wmin w1 w2)))\ eq2 min_less_iff_disj)+ show "wmin w1 w2 \\<^sub>E min r1 r2" using res1 res2 by auto qed lemma min_repU1: assumes "w1 \\<^sub>L x" assumes "w2 \\<^sub>L y" shows "wmin w1 w2 \\<^sub>L x " using wmin_lemma assms repL_def by (meson min_le_iff_disj) lemma min_repU2: assumes "w1 \\<^sub>L y" assumes "w2 \\<^sub>L x" shows "wmin w1 w2 \\<^sub>L x" using wmin_lemma assms repL_def by (meson min_le_iff_disj) subsection \Multiplication lower bound\ text\Multiplication lower bound\ fun tl :: "word \ word \ word \ word \ word" where "tl w1l w1u w2l w2u = wmin (wmin (wtimes w1l w2l) (wtimes w1u w2l)) (wmin (wtimes w1l w2u) (wtimes w1u w2u))" text\Correctness of multiplication lower bound\ lemma tl_lemma: assumes u1:"u\<^sub>1 \\<^sub>U (r1::real)" assumes u2:"u\<^sub>2 \\<^sub>U (r2::real)" assumes l1:"l\<^sub>1 \\<^sub>L (r1::real)" assumes l2:"l\<^sub>2 \\<^sub>L (r2::real)" shows "tl l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>L (r1 * r2)" proof - obtain rl1 rl2 ru1 ru2 :: real where gru1:"ru1 \ r1" and gru2:"ru2 \ r2" and grl1:"rl1 \ r1" and grl2:"rl2 \ r2" and eru1:"u\<^sub>1 \\<^sub>E ru1" and eru2:"u\<^sub>2 \\<^sub>E ru2" and erl1:"l\<^sub>1 \\<^sub>E rl1" and erl2:"l\<^sub>2 \\<^sub>E rl2" using u1 u2 l1 l2 unfolding repU_def repL_def by auto have timesuu:"wtimes u\<^sub>1 u\<^sub>2 \\<^sub>E ru1 * ru2" using wtimes_exact[OF eru1 eru2] by auto have timesul:"wtimes u\<^sub>1 l\<^sub>2 \\<^sub>E ru1 * rl2" using wtimes_exact[OF eru1 erl2] by auto have timeslu:"wtimes l\<^sub>1 u\<^sub>2 \\<^sub>E rl1 * ru2" using wtimes_exact[OF erl1 eru2] by auto have timesll:"wtimes l\<^sub>1 l\<^sub>2 \\<^sub>E rl1 * rl2" using wtimes_exact[OF erl1 erl2] by auto have maxt12:"wmin (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2) \\<^sub>E min (rl1 * rl2) (ru1 * rl2)" by (rule wmin_lemma[OF timesll timesul]) have maxt34:"wmin (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2) \\<^sub>E min (rl1 * ru2) (ru1 * ru2)" by (rule wmin_lemma[OF timeslu timesuu]) have bigMax:"wmin (wmin (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2)) (wmin (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2)) \\<^sub>E min (min(rl1 * rl2) (ru1 * rl2)) (min (rl1 * ru2) (ru1 * ru2))" by (rule wmin_lemma[OF maxt12 maxt34]) obtain maxt12val :: real where maxU12:"wmin (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2) \\<^sub>L min (rl1 * rl2) (ru1 * rl2)" using maxt12 unfolding repL_def by blast obtain maxt34val :: real where maxU34:"wmin (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2) \\<^sub>L min (rl1 * ru2) (ru1 * ru2)" using maxt34 unfolding repL_def by blast obtain bigMaxU:"wmin (wmin (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2)) (wmin (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2)) \\<^sub>L min (min (rl1 * rl2) (ru1 * rl2)) (min (rl1 * ru2) (ru1 * ru2))" using bigMax unfolding repL_def by blast have ivl1:"rl1 \ ru1" using grl1 gru1 by auto have ivl2:"rl2 \ ru2" using grl2 gru2 by auto let ?thesis = "tl l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>L r1 * r2" show ?thesis using ivl1 ivl2 proof(cases rule: case_ivl_zero) case ZeroZero assume "rl1 \ 0 \ 0 \ ru1 \ rl2 \ 0 \ 0 \ ru2" then have geq1:"ru1 \ 0" and geq2:"ru2 \ 0" and geq3:"rl1 \ 0" and geq4:"rl2 \ 0" by auto consider "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 have g1:"rl1 * ru2 \ 0" using "1" geq1 geq2 geq3 geq4 grl2 gru2 mult_le_0_iff by blast have g2:"0 \ r1 * r2" using "1" geq1 geq2 grl1 grl2 gru1 gru2 by (simp) from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show ?thesis using up eru1 eru2 erl1 erl2 min_repU1 min_repU2 repL_def repU_def timeslu tl.simps wmin.elims by (metis bigMax min_le_iff_disj) next case 2 have g1:"rl1 * ru2 \ rl1 * r2" using "2" geq1 geq2 grl2 gru2 by (metis mult_le_cancel_left geq3 leD) have g2:"rl1 * r2 \ r1 * r2" using "2" geq1 geq2 grl2 gru2 by (simp add: mult_right_mono grl1) from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show ?thesis by (metis up maxU12 min_repU2 repL_def tl.simps min.coboundedI1 maxt34) next case 3 have g1:"ru1 * rl2 \ ru1 * r2" using "3" geq1 geq2 grl2 gru2 by (simp add: mult_left_mono) have g2:"ru1 * r2 \ r1 * r2" using "3" geq1 geq2 grl1 grl2 gru1 gru2 mult_minus_right mult_right_mono by (simp add: mult_right_mono_neg) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 maxt34 timesul by (metis repL_def tl.simps) next case 4 have g1:"ru1 * rl2 \ 0" using "4" geq1 geq2 grl1 grl2 gru1 gru2 \rl1 \ 0 \ 0 \ ru1 \ rl2 \ 0 \ 0 \ ru2\ mult_less_0_iff less_eq_real_def not_less by auto have g2:"0 \ r1 * r2" using "4" geq1 geq2 grl1 grl2 gru1 gru2 by (metis mult_less_0_iff not_less) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis by (metis up maxU12 maxU34 wmin.elims min_repU1 min_repU2 repL_def timesul tl.simps) qed next case ZeroPos assume bounds:"rl1 \ 0 \ 0 \ ru1 \ 0 \ rl2" have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast consider "r1 \ 0" | "r1 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 have g1:"rl1 * rl2 \ 0" using "1" r2 bounds grl1 grl2 gru1 gru2 by (simp add: mult_le_0_iff) have g2:"0 \ r1 * r2" using "1" r2 bounds grl1 grl2 gru1 gru2 by (simp) from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis by (metis repL_def timesll tl.simps up maxU12 maxU34 wmin.elims min_repU2 min_repU1) next case 2 have bound:"ru2 \ 0" using "2" r2 bounds grl1 grl2 gru1 gru2 dual_order.trans by auto then have g1:"rl1 * ru2 \ rl1 * r2" using "2" r2 bounds grl1 grl2 gru1 gru2 mult_le_cancel_left by fastforce have g2:"rl1 * r2 \ r1 * r2" using "2" r2 bounds grl1 grl2 gru1 gru2 mult_le_0_iff mult_le_cancel_right by fastforce from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show ?thesis by (metis up maxU12 wmin.elims min_repU2 min.coboundedI1 maxt34 repL_def tl.simps) qed next case ZeroNeg assume bounds:"rl1 \ 0 \ 0 \ ru1 \ ru2 \ 0" have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast consider (Pos) "r1 \ 0" | (Neg) "r1 \ 0" using le_cases by auto then show ?thesis proof (cases) case Pos have bound:"rl2 \ 0" using Pos r2 bounds grl1 grl2 gru1 gru2 dual_order.trans by auto then have g1:"ru1 * rl2 \ ru1 * r2" using Pos bounds grl1 grl2 gru1 gru2 mult_le_cancel_left by fastforce have p1:"\a::real. (0 \ - a) = (a \ 0)" by(auto) have p2:"\a b::real. (- a \ - b) = (b \ a)" by auto have g2:"ru1 * r2 \ r1 * r2" using Pos r2 bounds grl1 grl2 gru1 gru2 p1 p2 by (simp add: mult_right_mono_neg) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis by (metis up maxU12 maxU34 wmin.elims min_repU2 min_repU1 repL_def timesul tl.simps) next case Neg have g1:"ru1 * ru2 \ 0" using Neg r2 bounds grl1 grl2 gru1 gru2 mult_le_0_iff by blast have g2:"0 \ r1 * r2" using Neg r2 zero_le_mult_iff by blast from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 min.coboundedI1 min.commute maxt34 by (metis repL_def tl.simps) qed next case PosZero assume bounds:"0 \ rl1 \ rl2 \ 0 \ 0 \ ru2" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast have bound:"0 \ ru1" using r1 bounds grl1 grl2 gru1 gru2 dual_order.trans by auto consider "r2 \ 0" | "r2 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 have g1:"rl1 * rl2 \ 0" using r1 "1" bounds grl1 grl2 gru1 gru2 mult_le_0_iff by blast have g2:"0 \ r1 * r2" using r1 "1" bounds grl1 grl2 gru1 gru2 zero_le_mult_iff by blast from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims min_repU2 min_repU1 min.coboundedI1 min.commute maxt12 maxt34 repL_def timesll tl.simps by metis next case 2 have g1:"ru1 * rl2 \ ru1 * r2" using r1 "2" bounds bound grl1 grl2 gru1 gru2 using mult_left_mono by blast have g2:"ru1 * r2 \ r1 * r2" using r1 "2" bounds bound grl2 gru2 by (metis mult_left_mono_neg gru1 mult.commute) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 maxt34 by (metis repL_def timesul tl.simps) qed next case NegZero assume bounds:"ru1 \ 0 \ rl2 \ 0 \ 0 \ ru2" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast have bound:"rl1 \ 0" using r1 bounds grl1 grl2 gru1 gru2 dual_order.trans by auto consider "r2 \ 0" | "r2 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 assume r2:"r2 \ 0" have g1:"rl1 * ru2 \ rl1 * r2" using r1 r2 bounds bound grl1 grl2 gru1 gru2 by (metis mult_le_cancel_left leD) have g2:"rl1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_right_mono by (simp add: mult_le_0_iff) from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU min_repU2 min_repU1 min.coboundedI1 maxt34 by (metis min_repU2 repL_def tl.simps) next case 2 assume r2:"r2 \ 0" have lower:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have g1:"ru1 * ru2 \ 0" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_le_0_iff by blast have g2:"0 \ r1 * r2" using r1 r2 by (simp add: zero_le_mult_iff) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 min.coboundedI1 min.commute maxt34 by (metis repL_def tl.simps) qed next case NegNeg assume bounds:"ru1 \ 0 \ ru2 \ 0" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast have lower1:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have lower2:"rl2 \ 0" using bounds dual_order.trans grl2 r2 by blast have g1:"ru1 * ru2 \ ru1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using not_less mult_le_cancel_left by metis have g2:"ru1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_le_cancel_left mult.commute not_le lower1 lower2 by metis from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 min.coboundedI1 min.commute maxt34 by (metis repL_def tl.simps) next case NegPos assume bounds:"ru1 \ 0 \ 0 \ rl2" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast have lower1:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have lower2:"rl2 \ 0" using bounds by auto have upper1:"ru1 \ 0" using bounds by auto have upper2:"ru2 \ 0" using bounds dual_order.trans gru2 r2 by blast have g1:"rl1 * ru2 \ rl1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 less_le_trans upper1 lower2 by (metis mult_le_cancel_left not_less) have g2:"rl1 * r2 \ r1 * r2" using r1 upper1 r2 mult_right_mono mult_le_0_iff grl1 by blast from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 maxt12 maxt34 by (metis repL_def timeslu tl.simps) next case PosNeg assume bounds:"0 \ rl1 \ ru2 \ 0" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast have lower1:"rl1 \ 0" using bounds by auto have lower2:"rl2 \ 0" using dual_order.trans grl2 r2 by blast have upper1:"ru1 \ 0" using dual_order.trans gru1 u1 r1 by blast have upper2:"ru2 \ 0" using bounds by auto have g1:"ru1 * rl2 \ ru1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_left_mono less_le_trans not_less by metis have g2:"ru1 * r2 \ r1 * r2" using r1 lower1 r2 not_less gru2 gru1 grl1 grl2 by (metis mult_le_cancel_left mult.commute) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show "tl l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>L r1 * r2" using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 by (metis repL_def timesul tl.simps) next case PosPos assume bounds:"0 \ rl1 \ 0 \ rl2" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast have lower1:"rl1 \ 0" using bounds by auto have lower2:"rl2 \ 0" using bounds by auto have upper1:"ru1 \ 0" using dual_order.trans gru1 u1 r1 by blast have upper2:"ru2 \ 0" using dual_order.trans gru2 u2 r2 bounds by blast have g1:"rl1 * rl2 \ rl1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_left_mono using leD leI less_le_trans by auto have g2:"rl1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_right_mono by blast from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU min_repU2 min_repU1 min.coboundedI1 maxt12 maxt34 by (metis repL_def tl.simps) qed qed text\Most significant bit only changes under successor when all other bits are 1\ lemma msb_succ: fixes w :: "32 Word.word" assumes neq1:"uint w \ 0xFFFFFFFF" assumes neq2:"uint w \ 0x7FFFFFFF" shows "msb (w + 1) = msb w" proof - have "w \ 0xFFFFFFFF" using neq1 by auto then have neqneg1:"w \ -1" by auto have "w \ 0x7FFFFFFF" using neq2 by auto then have neqneg2:"w \ (2^31)-1" by auto show ?thesis using neq1 neq2 unfolding msb_big using Word_Lemmas.word_le_make_less[of "w + 1" "0x80000000"] Word_Lemmas.word_le_make_less[of "w " "0x80000000"] neqneg1 neqneg2 by auto qed text\Negation commutes with msb except at edge cases\ lemma msb_non_min: fixes w :: "32 Word.word" assumes neq1:"uint w \ 0" assumes neq2:"uint w \ ((2^(len_of (TYPE(31)))))" shows "msb (uminus w) = HOL.Not(msb(w))" proof - have fact1:"uminus w = word_succ (~~ w)" by (rule twos_complement) have fact2:"msb (~~w) = HOL.Not(msb w)" using Word.word_ops_msb[of w] by auto have neqneg1:"w \ 0" using neq1 by auto have not_undef:"w \ 0x80000000" using neq2 by auto then have neqneg2:"w \ (2^31)" by auto have "~~ 0xFFFFFFFF = (0::word)" by auto then have "(~~ w) \ 0xFFFFFFFF" using neqneg1 neqneg2 Word.word_bool_alg.double_compl[of w] by metis then have uintNeq1:"uint (~~ w) \ 0xFFFFFFFF" using uint_distinct[of "~~w" "0xFFFFFFFF"] by auto have "~~ (0x7FFFFFFF::word) = (2 ^ 31::word)" by simp then have "(~~ w) \ 0x7FFFFFFF" using neqneg1 neqneg2 Word.word_bool_alg.double_compl[of w] by metis then have uintNeq2:" uint (~~ w) \ 0x7FFFFFFF" using uint_distinct[of "~~w" "0x7FFFFFFF"] by auto have fact3:"msb ((~~w) + 1) = msb (~~w)" apply(rule msb_succ[of "~~w"]) using neq1 neq2 uintNeq1 uintNeq2 by auto show "msb (uminus w) = HOL.Not(msb(w))" using fact1 fact2 fact3 by (simp add: word_succ_p1) qed text\Only 0x80000000 preserves msb=1 under negation\ lemma msb_min_neg: fixes w::"word" assumes msb1:"msb (- w)" assumes msb2:"msb w" shows "uint w = ((2^(len_of (TYPE(31)))))" - proof - - have neq1:"uint w \ 0" using msb2 word_msb_0 uint_0_iff by force - show ?thesis using msb1 msb2 msb_non_min[OF neq1] - by auto - qed - +proof (rule ccontr) + from \msb w\ have \w \ 0\ + using word_msb_0 by auto + then have \uint w \ 0\ + by transfer simp + moreover assume \uint w \ 2 ^ LENGTH(31)\ + ultimately have \msb (- w) \ \ msb w\ + by (rule msb_non_min) + with assms show False + by simp +qed text\Only 0x00000000 preserves msb=0 under negation\ lemma msb_zero: fixes w::"word" assumes msb1:"\ msb (- w)" assumes msb2:"\ msb w" shows "uint w = 0" proof - have neq:"w \ ((2 ^ len_of TYPE(31))::word)" using msb1 msb2 by auto have eq:"uint ((2 ^ len_of TYPE(31))::word) = 2 ^ len_of TYPE(31)" by auto then have neq:"uint w \ uint ((2 ^ len_of TYPE(31))::word)" using uint_distinct[of w "2^len_of TYPE(31)"] neq eq by auto show ?thesis using msb1 msb2 minus_zero msb_non_min[of w] neq by force qed text\Finite numbers alternate msb under negation\ lemma msb_pos: fixes w::"word" assumes msb1:"msb (- w)" assumes msb2:"\ msb w" shows "uint w \ {1 .. (2^((len_of TYPE(32)) - 1))-1}" proof - have main: "w \ {1 .. (2^((len_of TYPE(32)) - 1))-1}" using msb1 msb2 apply(clarsimp) unfolding word_msb_sint apply(rule conjI) apply (metis neg_equal_0_iff_equal not_le word_less_1) proof - have imp:"w \ 0x80000000 \ False" proof - assume geq:"w \ 0x80000000" then have "msb w" using Word_Lemmas.msb_big[of w] by auto then show False using msb2 by auto qed have mylem:"\w1 w2::word. uint w1 \ uint w2 \ w1 \ w2" subgoal for w1 w2 by (simp add: word_le_def) done have mylem2:"\w1 w2::word. w1 > w2 \ uint w1 > uint w2" subgoal for w1 w2 by (simp add: word_less_def) done have gr_to_geq:"w > 0x7FFFFFFF \ w \ 0x80000000" apply(rule mylem) using mylem2[of "0x7FFFFFFF" "w"] by auto have taut:"w \ 0x7FFFFFFF \ w > 0x7FFFFFFF" by auto then show "w \ 0x7FFFFFFF" using imp taut gr_to_geq by auto qed have set_eq:"(uint ` (({1..(minus(2 ^ (minus(len_of TYPE(32)) 1)) 1)})::word set)) = ({1..minus(2 ^ (minus (len_of TYPE(32)) 1)) 1}::int set)" apply(auto simp add: word_le_def) subgoal for xa proof - assume lower:"1 \ xa" and upper:"xa \ 2147483647" then have in_range:"xa \ {0 .. 2^32-1}" by auto then have "xa \ range (uint::word \ int)" unfolding Word.word_uint.Rep_range Word.uints_num by auto then obtain w::word where xaw:"xa = uint w" by auto then have "w \ {1..0x7FFFFFFF} " using lower upper apply(clarsimp, auto) by (auto simp add: word_le_def) then show ?thesis using uint_distinct uint_distinct main image_eqI word_le_def xaw by blast qed done then show "uint w \ {1..2 ^ (len_of TYPE(32) - 1) - 1}" using uint_distinct uint_distinct main image_eqI by blast qed lemma msb_neg: fixes w::"word" assumes msb1:"\ msb (- w)" assumes msb2:"msb w" shows "uint w \ {2^((len_of TYPE(32) - 1))+1 .. 2^((len_of TYPE(32)))-1}" proof - have mylem:"\w1 w2::word. uint w1 \ uint w2 \ w1 \ w2" by (simp add: word_le_def) have mylem2:"\w1 w2::word. w1 > w2 \ uint w1 > uint w2" by (simp add: word_less_def) have gr_to_geq:"w > 0x80000000 \ w \ 0x80000001" apply(rule mylem) using mylem2[of "0x80000000" "w"] by auto have taut:"w \ 0x80000000 \ 0x80000000 < w" by auto have imp:"w \ 0x80000000 \ False" proof - assume geq:"w \ 0x80000000" then have "(msb (-w))" using Word_Lemmas.msb_big[of "-w"] Word_Lemmas.msb_big[of "w"] by (simp add: msb2) then show False using msb1 by auto qed have main: "w \ {2^((len_of TYPE(32)) - 1)+1 .. 2^((len_of TYPE(32)))-1}" using msb1 msb2 apply(clarsimp) unfolding word_msb_sint proof - show "0x80000001 \ w" using imp taut gr_to_geq by auto qed have set_eq:"(uint ` (({2^((len_of TYPE(32) - 1))+1 .. 2^((len_of TYPE(32)))-1})::word set)) = {2^((len_of TYPE(32) - 1))+1 .. 2^((len_of TYPE(32)))-1}" apply(auto) subgoal for xa by (simp add: word_le_def) subgoal for w using uint_lt [of w] by simp subgoal for xa proof - assume lower:"2147483649 \ xa" and upper:"xa \ 4294967295" then have in_range:"xa \ {0x80000000 .. 0xFFFFFFFF}" by auto then have "xa \ range (uint::word \ int)" unfolding Word.word_uint.Rep_range Word.uints_num by auto then obtain w::word where xaw:"xa = uint w" by auto then have the_in:"w \ {0x80000001 .. 0xFFFFFFFF} " using lower upper by (auto simp add: word_le_def) have the_eq:"(0xFFFFFFFF::word) = -1" by auto from the_in the_eq have "w \ {0x80000001 .. -1}" by auto then show ?thesis using uint_distinct uint_distinct main image_eqI word_le_def xaw by blast qed done then show "uint w \ {2^((len_of TYPE(32)) - 1)+1 .. 2^((len_of TYPE(32)))-1}" using uint_distinct uint_distinct main image_eqI by blast qed text\2s-complement commutes with negation except edge cases\ lemma sint_neg_hom: fixes w :: "32 Word.word" shows "uint w \ ((2^(len_of (TYPE(31))))) \ (sint(-w) = -(sint w))" unfolding Word_Bitwise.word_sint_msb_eq apply auto subgoal using msb_min_neg by auto prefer 3 subgoal using msb_zero[of w] by (simp add: msb_zero) proof - assume msb1:"msb (- w)" assume msb2:"\ msb w" have "uint w \ {1 .. (2^((len_of TYPE(32)) - 1))-1}" using msb_pos[OF msb1 msb2] by auto then have bound:"uint w \ {1 .. 0x7FFFFFFF}" by auto have size:"size (w::32 Word.word) = 32" using Word.word_size[of w] by auto have lem:"\x::int. \n::nat. x \ {1..(2^n)-1} \ ((- x) mod (2^n)) - (2^n) = - x" subgoal for x n apply(cases "x mod 2^n = 0") by(auto simp add: Divides.zmod_zminus1_eq_if[of x "2^n"]) done have lem_rule:"uint w \ {1..2 ^ 32 - 1} \ (- uint w mod 4294967296) - 4294967296 = - uint w" using lem[of "uint w" 32] by auto have almost:"- uint w mod 4294967296 - 4294967296 = - uint w" apply(rule lem_rule) using bound by auto show "uint (- w) - 2 ^ size (- w) = - uint w" using bound unfolding Word.uint_word_ariths word_size_neg by (auto simp add: size almost) next assume neq:"uint w \ 0x80000000" assume msb1:"\ msb (- w)" assume msb2:"msb w" have bound:"uint w \ {0x80000001.. 0xFFFFFFFF}" using msb1 msb2 msb_neg by auto have size:"size (w::32 Word.word) = 32" using Word.word_size[of w] by auto have lem:"\x::int. \n::nat. x \ {1..(2^n)-1} \ (-x mod (2^n)) = (2^n) - x" subgoal for x n apply(auto) apply(cases "x mod 2^n = 0") by (simp add: Divides.zmod_zminus1_eq_if[of x "2^n"])+ done from bound have wLeq: "uint w \ 4294967295" and wGeq: "2147483649 \ uint w" by auto from wLeq have wLeq':"uint w \ 4294967296" by fastforce have f3: "(0 \ 4294967296 + - 1 * uint w + - 1 * ((4294967296 + - 1 * uint w) mod 4294967296)) = (uint w + (4294967296 + - 1 * uint w) mod 4294967296 \ 4294967296)" by auto have f4: "(0 \ 4294967296 + - 1 * uint w) = (uint w \ 4294967296)" by auto have f5: "\i ia. \ (0::int) \ i \ 0 \ i + - 1 * (i mod ia)" by (simp add: zmod_le_nonneg_dividend) then have f6: "uint w + (4294967296 + - 1 * uint w) mod 4294967296 \ 4294967296" using f4 f3 wLeq' by blast have f7: "4294967296 + - 1 * uint w + - 4294967296 = - 1 * uint w" by auto have f8: "- (1::int) * 4294967296 = - 4294967296" by auto have f9: "(0 \ - 1 * uint w) = (uint w \ 0)" by auto have f10: "(4294967296 + -1 * uint w + -1 * ((4294967296 + -1 * uint w) mod 4294967296) \ 0) = (4294967296 \ uint w + (4294967296 + - 1 * uint w) mod 4294967296)" by auto have f11: "\ 4294967296 \ (0::int)" by auto have f12: "\x0. ((0::int) < x0) = (\ x0 \ 0)" by auto have f13: "\x0 x1. ((x1::int) < x0) = (\ 0 \ x1 + - 1 * x0)" by auto have f14: "\x0 x1. ((x1::int) \ x1 mod x0) = (x1 + - 1 * (x1 mod x0) \ 0)" by auto have "\ uint w \ 0" using wGeq by fastforce then have "4294967296 \ uint w + (4294967296 + - 1 * uint w) mod 4294967296" using f14 f13 f12 f11 f10 f9 f8 f7 by (metis (no_types) int_mod_ge) then show "uint (- w) = 2 ^ size w - uint w" using f6 unfolding Word.uint_word_ariths by (auto simp add: size f4) qed text\2s-complement encoding is injective\ lemma sint_dist: fixes x y ::word assumes "x \ y" shows "sint x \ sint y" by (simp add: assms) subsection\Negation\ fun wneg :: "word \ word" where "wneg w = (if w = NEG_INF then POS_INF else if w = POS_INF then NEG_INF else -w)" text\word negation is correct\ lemma wneg_lemma: assumes eq:"w \\<^sub>E (r::real)" shows "wneg w \\<^sub>E -r" apply(rule repe.cases[OF eq]) apply(auto intro!: repNEG_INF repPOS_INF simp add: repe.simps)[2] subgoal for ra proof - assume eq:"w = ra" assume i:"r = (real_of_int (sint ra))" assume bounda:" (real_of_int (sint ra)) < (real_of_int (sint POS_INF))" assume boundb:" (real_of_int (sint NEG_INF)) < (real_of_int (sint ra))" have raNeq:"ra \ 2147483647" using sint_range[OF bounda boundb] by (auto) have raNeqUndef:"ra \ 2147483648" using int_not_undef[OF bounda boundb] by (auto) have "uint ra \ uint ((2 ^ len_of TYPE(31))::word)" apply (rule uint_distinct) using raNeqUndef by auto then have raNeqUndefUint:"uint ra \ ((2 ^ len_of TYPE(31)))" by auto have res1:"wneg w \\<^sub>E (real_of_int (sint (wneg w)))" apply (rule repINT) using sint_range[OF bounda boundb] sint_neg_hom[of ra, OF raNeqUndefUint] raNeq raNeqUndefUint raNeqUndef eq by(auto) have res2:"- r = (real_of_int (sint (wneg w)))" using eq bounda boundb i sint_neg_hom[of ra, OF raNeqUndefUint] raNeq raNeqUndef eq by (auto) show ?thesis using res1 res2 by auto qed done subsection\Comparison\ fun wgreater :: "word \ word \ bool" where "wgreater w1 w2 = (sint w1 > sint w2)" lemma neg_less_contra:"\x. Suc x < - (Suc x) \ False" by auto text\Comparison < is correct\ lemma wgreater_lemma:"w1 \\<^sub>L (r1::real) \ w2 \\<^sub>U r2 \ wgreater w1 w2 \ r1 > r2" proof (auto simp add: repU_def repL_def) fix r'\<^sub>1 r'\<^sub>2 assume sint_le:"sint w1 > sint w2" then have sless:"(w2 1 \ r1" assume r2_leq:"r2 \ r'\<^sub>2" assume wr1:"w1 \\<^sub>E r'\<^sub>1" assume wr2:"w2 \\<^sub>E r'\<^sub>2" have greater:"r'\<^sub>1 > r'\<^sub>2" using wr1 wr2 apply(auto simp add: repe.simps) prefer 4 using sless sint_le by (auto simp add: less_le_trans not_le) show "r1 > r2" using r1_leq r2_leq greater by auto qed text\Comparison $\geq$ of words\ fun wgeq :: "word \ word \ bool" where "wgeq w1 w2 = ((\ ((w2 = NEG_INF \ w1 = NEG_INF) \(w2 = POS_INF \ w1 = POS_INF))) \ (sint w2 \ sint w1))" text\Comparison $\geq$ of words is correct\ lemma wgeq_lemma:"w1 \\<^sub>L r1 \ w2 \\<^sub>U (r2::real) \ wgeq w1 w2 \ r1 \ r2" proof (unfold wgeq.simps) assume assms:"\ (w2 = NEG_INF \ w1 = NEG_INF \ w2 = POS_INF \ w1 = POS_INF) \ sint w2 \ sint w1" assume a1:"w1 \\<^sub>L r1" and a2:"w2 \\<^sub>U (r2::real)" from assms have sint_le:"sint w2 \ sint w1" by auto then have sless:"(w2 <=s w1)" using word_sless_alt word_sle_def by auto obtain r'\<^sub>1 r'\<^sub>2 where r1_leq:"r'\<^sub>1 \ r1" and r2_leq:"r2 \ r'\<^sub>2" and wr1:"w1 \\<^sub>E r'\<^sub>1" and wr2:"w2 \\<^sub>E r'\<^sub>2" using a1 a2 unfolding repU_def repL_def by auto from assms have check1:"\ (w1 = NEG_INF \ w2 = NEG_INF)" by auto from assms have check2:"\ (w1 = POS_INF \ w2 = POS_INF)" by auto have less:"r'\<^sub>2 \ r'\<^sub>1" using sless sint_le check1 check2 repe.simps wr2 wr1 by(auto simp add: repe.simps) show "r1 \ r2" using r1_leq r2_leq less by auto qed subsection\Absolute value\ text\Absolute value of word\ fun wabs :: "word \ word" where "wabs l1 = (wmax l1 (wneg l1))" text\Correctness of wmax\ lemma wabs_lemma: assumes eq:"w \\<^sub>E (r::real)" shows "wabs w \\<^sub>E (abs r)" proof - have w:"wmax w (wneg w) \\<^sub>E max r (-r)" by (rule wmax_lemma[OF eq wneg_lemma[OF eq]]) have r:"max r (-r) = abs r" by auto from w r show ?thesis by auto qed end \ No newline at end of file diff --git a/thys/Iptables_Semantics/Common/Word_Upto.thy b/thys/Iptables_Semantics/Common/Word_Upto.thy --- a/thys/Iptables_Semantics/Common/Word_Upto.thy +++ b/thys/Iptables_Semantics/Common/Word_Upto.thy @@ -1,203 +1,203 @@ section\Word Upto\ theory Word_Upto imports Main IP_Addresses.Hs_Compat Word_Lib.Word_Lemmas begin text\Enumerate a range of machine words.\ text\enumerate from the back (inefficient)\ -function word_upto :: "'a word \ 'a word \ ('a::len0) word list" where +function word_upto :: "'a word \ 'a word \ ('a::len) word list" where "word_upto a b = (if a = b then [a] else word_upto a (b - 1) @ [b])" by pat_completeness auto (*by the way: does not terminate practically if b < a; will terminate after it reaches the word wrap-around!*) termination word_upto apply(relation "measure (unat \ uncurry (-) \ prod.swap)") apply(rule wf_measure; fail) apply(simp) apply(subgoal_tac "unat (b - a - 1) < unat (b - a)") apply(simp add: diff_right_commute; fail) apply(rule measure_unat) apply auto done declare word_upto.simps[simp del] text\enumerate from the front (more inefficient)\ -function word_upto' :: "'a word \ 'a word \ ('a::len0) word list" where +function word_upto' :: "'a word \ 'a word \ ('a::len) word list" where "word_upto' a b = (if a = b then [a] else a # word_upto' (a + 1) b)" by pat_completeness auto termination word_upto' apply(relation "measure (\ (a, b). unat (b - a))") apply(rule wf_measure; fail) apply(simp) apply(subgoal_tac "unat (b - a - 1) < unat (b - a)") apply (simp add: diff_diff_add; fail) apply(rule measure_unat) apply auto done declare word_upto'.simps[simp del] lemma word_upto_cons_front[code]: "word_upto a b = word_upto' a b" proof(induction a b rule:word_upto'.induct) case (1 a b) have hlp1: "a \ b \ a # word_upto (a + 1) b = word_upto a b" apply(induction a b rule:word_upto.induct) apply simp apply(subst(1) word_upto.simps) apply(simp) apply safe apply(subst(1) word_upto.simps) apply (simp) apply(subst(1) word_upto.simps) apply (simp; fail) apply(case_tac "a \ b - 1") apply(simp) apply (metis Cons_eq_appendI word_upto.simps) apply(simp) done from 1[symmetric] show ?case apply(cases "a = b") subgoal apply(subst word_upto.simps) apply(subst word_upto'.simps) by(simp) apply(subst word_upto'.simps) by(simp add: hlp1) qed (* Most of the lemmas I show about word_upto hold without a \ b, but I don't need that right now and it's giving me a headache *) lemma word_upto_set_eq: "a \ b \ x \ set (word_upto a b) \ a \ x \ x \ b" proof show "a \ b \ x \ set (word_upto a b) \ a \ x \ x \ b" apply(induction a b rule: word_upto.induct) apply(case_tac "a = b") apply(subst(asm) word_upto.simps) apply(simp; fail) apply(subst(asm) word_upto.simps) apply(simp) apply(erule disjE) apply(simp; fail) proof(goal_cases) case (1 a b) from 1(2-3) have "b \ 0" by force from 1(2,3) have "a \ b - 1" by (metis \b \ 0\ eq_iff le_step_down_nat unat_arith_simps(1) unat_minus_one) from 1(1)[OF this 1(4)] show ?case by (metis dual_order.trans 1(2,3) less_imp_le measure_unat word_le_0_iff word_le_nat_alt) qed next show "a \ x \ x \ b \ x \ set (word_upto a b)" apply(induction a b rule: word_upto.induct) apply(case_tac "a = b") apply(subst word_upto.simps) apply(simp; force) apply(subst word_upto.simps) apply(simp) apply(case_tac "x = b") apply(simp;fail) proof(goal_cases) case (1 a b) from 1(2-4) have "b \ 0" by force from 1(2,4) have "x \ b - 1" by (metis \b \ 0\ dual_order.antisym le_step_down_nat unat_minus_one word_le_nat_alt) from 1(1) this show ?case by simp qed qed lemma word_upto_distinct_hlp: "a \ b \ a \ b \ b \ set (word_upto a (b - 1))" apply(rule ccontr, unfold not_not) apply(subgoal_tac "a \ b - 1") apply(drule iffD1[OF word_upto_set_eq[of a "b -1" b]]) apply(simp add: word_upto.simps; fail) apply(subgoal_tac "b \ 0") apply(meson leD measure_unat word_le_nat_alt) apply(blast intro: iffD1[OF word_le_0_iff]) apply(metis antisym_conv le_step_down_nat unat_minus_one word_le_0_iff word_le_nat_alt) done lemma distinct_word_upto: "a \ b \ distinct (word_upto a b)" apply(induction a b rule: word_upto.induct) apply(case_tac "a = b") apply(subst word_upto.simps) apply(simp; force) apply(subst word_upto.simps) apply(case_tac "a \ b - 1") apply(simp) apply(rule word_upto_distinct_hlp; simp; fail) apply(simp) apply(rule ccontr) apply(metis le_step_down_nat less_le not_le unat_minus_one word_le_nat_alt word_not_simps(1)) done lemma word_upto_eq_upto: "s \ e \ e \ unat (max_word :: 'l word) \ word_upto ((of_nat :: nat \ ('l :: len) word) s) (of_nat e) = map of_nat (upt s (Suc e))" proof(induction e) let ?mwon = "of_nat :: nat \ 'l word" let ?mmw = "max_word :: 'l word" case (Suc e) show ?case proof(cases "?mwon s = ?mwon (Suc e)") case True have "s = Suc e" using le_unat_uoi Suc.prems True by metis with True show ?thesis by(subst word_upto.simps) (simp) next case False hence le: "s \ e" using le_SucE Suc.prems by blast have lm: "e \ unat ?mmw" using Suc.prems by simp have sucm: "(of_nat :: nat \ ('l :: len) word) (Suc e) - 1 = of_nat e" using Suc.prems(2) by simp note mIH = Suc.IH[OF le lm] show ?thesis by(subst word_upto.simps) (simp add: False[simplified] Suc.prems mIH sucm) qed qed(simp add: word_upto.simps) lemma word_upto_alt: "(a :: ('l :: len) word) \ b \ word_upto a b = map of_nat (upt (unat a) (Suc (unat b)))" proof - let ?mmw = "max_word :: 'l word" assume le: "a \ b" hence nle: "unat a \ unat b" by(unat_arith) have lem: "unat b \ unat ?mmw" by (simp add: word_unat_less_le) note word_upto_eq_upto[OF nle lem, unfolded word_unat.Rep_inverse] thus "word_upto a b = map of_nat [unat a.. b then map of_nat (upt (unat a) (Suc (unat b))) else word_upto a b)" using word_upto_alt by metis lemma sorted_word_upto: fixes a b :: "('l :: len) word" assumes "a \ b" shows "sorted (word_upto a b)" using assms proof(induction b) fix b assume le_prem: "a \ 1 + b" assume nmax_prem: "1 + b \ 0" assume IH: "a \ b \ sorted (word_upto a b)" show "sorted (word_upto a (1 + b))" proof(cases "a = 1 + b") case True thus ?thesis by(simp add: word_upto.simps) next case False have fprem: "a \ b" using le_prem False by (metis add.commute antisym_conv1 plus_one_helper) note mIH = IH[OF this] show ?thesis by(subst word_upto.simps) (simp add: fprem lt1_neq0 nmax_prem word_le_plus_either word_upto_set_eq False sorted_append mIH) qed qed(simp add: word_upto.simps) end diff --git a/thys/Native_Word/Code_Target_Word_Base.thy b/thys/Native_Word/Code_Target_Word_Base.thy --- a/thys/Native_Word/Code_Target_Word_Base.thy +++ b/thys/Native_Word/Code_Target_Word_Base.thy @@ -1,424 +1,424 @@ (* Title: Code_Target_Word_Base.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Common base for target language implementations of word types\ theory Code_Target_Word_Base imports "HOL-Word.Word" Bits_Integer begin text \More lemmas\ lemma nat_div_eq_Suc_0_iff: "n div m = Suc 0 \ n \ m \ n < 2 * m" apply (auto simp add: sdl) using not_less apply fastforce apply (metis One_nat_def Suc_1 div_eq_0_iff lessI neq0_conv td_gal_lt) done lemma Suc_0_lt_2p_len_of: "Suc 0 < 2 ^ LENGTH('a :: len)" by (metis One_nat_def len_gt_0 lessI numeral_2_eq_2 one_less_power) lemma div_half_nat: fixes x y :: nat assumes "y \ 0" shows "(x div y, x mod y) = (let q = 2 * (x div 2 div y); r = x - q * y in if y \ r then (q + 1, r - y) else (q, r))" proof - let ?q = "2 * (x div 2 div y)" have q: "?q = x div y - x div y mod 2" by(metis div_mult2_eq mult.commute minus_mod_eq_mult_div [symmetric]) let ?r = "x - ?q * y" have r: "?r = x mod y + x div y mod 2 * y" by(simp add: q diff_mult_distrib minus_mod_eq_div_mult [symmetric])(metis diff_diff_cancel mod_less_eq_dividend mod_mult2_eq add.commute mult.commute) show ?thesis proof(cases "y \ x - ?q * y") case True with assms q have "x div y mod 2 \ 0" unfolding r by (metis Nat.add_0_right diff_0_eq_0 diff_Suc_1 le_div_geq mod2_gr_0 mod_div_trivial mult_0 neq0_conv numeral_1_eq_Suc_0 numerals(1)) hence "x div y = ?q + 1" unfolding q by simp moreover hence "x mod y = ?r - y" by simp(metis minus_div_mult_eq_mod [symmetric] diff_commute diff_diff_left mult_Suc) ultimately show ?thesis using True by(simp add: Let_def) next case False hence "x div y mod 2 = 0" unfolding r by(simp add: not_le)(metis Nat.add_0_right assms div_less div_mult_self2 mod_div_trivial mult.commute) hence "x div y = ?q" unfolding q by simp moreover hence "x mod y = ?r" by (metis minus_div_mult_eq_mod [symmetric]) ultimately show ?thesis using False by(simp add: Let_def) qed qed lemma unat_p2: "n < LENGTH('a :: len) \ unat (2 ^ n :: 'a word) = 2 ^ n" proof(induct n) case 0 thus ?case by simp next case (Suc n) then obtain n' where "LENGTH('a) = Suc n'" by(cases "LENGTH('a)") simp_all with Suc show ?case by (simp add: unat_word_ariths bintrunc_mod2p) qed lemma word_div_lt_eq_0: "x < y \ x div y = 0" for x :: "'a :: len word" by (simp add: word_eq_iff word_less_def word_test_bit_def uint_div) lemma word_div_eq_1_iff: "n div m = 1 \ n \ m \ unat n < 2 * unat (m :: 'a :: len word)" apply(simp only: word_arith_nat_defs word_le_nat_alt nat_div_eq_Suc_0_iff[symmetric]) apply(rule word_unat.Abs_inject) apply(simp only: unat_div[symmetric] word_unat.Rep) apply(simp add: unats_def Suc_0_lt_2p_len_of) done lemma div_half_word: fixes x y :: "'a :: len word" assumes "y \ 0" shows "(x div y, x mod y) = (let q = (x >> 1) div y << 1; r = x - q * y in if y \ r then (q + 1, r - y) else (q, r))" proof - obtain n where n: "x = of_nat n" "n < 2 ^ LENGTH('a)" by (cases x) obtain m where m: "y = of_nat m" "m < 2 ^ LENGTH('a)" by(cases y) let ?q = "(x >> 1) div y << 1" let ?q' = "2 * (n div 2 div m)" have "n div 2 div m < 2 ^ LENGTH('a)" using n by (metis of_nat_inverse unat_lt2p uno_simps(2)) hence q: "?q = of_nat ?q'" using n m apply (simp add: shiftr_def shiftr1_def uint_nat unat_of_nat shiftl_def shiftl1_def Bit_def word_div_def word_of_nat) apply (metis of_nat_inverse of_nat_numeral uno_simps(2) word_of_nat zdiv_int) done from assms have "m \ 0" using m by -(rule notI, simp) from n have "2 * (n div 2 div m) < 2 ^ LENGTH('a)" by(metis mult.commute div_mult2_eq minus_mod_eq_mult_div [symmetric] less_imp_diff_less of_nat_inverse unat_lt2p uno_simps(2)) moreover have "2 * (n div 2 div m) * m < 2 ^ LENGTH('a)" using n unfolding div_mult2_eq[symmetric] by(subst (2) mult.commute)(simp add: minus_mod_eq_div_mult [symmetric] diff_mult_distrib minus_mod_eq_mult_div [symmetric] div_mult2_eq) moreover have "2 * (n div 2 div m) * m \ n" by(metis div_mult2_eq div_mult_le mult.assoc mult.commute) ultimately have r: "x - ?q * y = of_nat (n - ?q' * m)" and "y \ x - ?q * y \ of_nat (n - ?q' * m) - y = of_nat (n - ?q' * m - m)" using n m unfolding q by (simp_all add: word_sub_wi word_mult_def uint_nat unat_of_nat of_nat_mult [symmetric] word_of_nat[symmetric] of_nat_diff word_le_nat_alt del: of_nat_mult) (metis diff_diff_left less_imp_diff_less of_nat_diff of_nat_inverse word_of_nat) thus ?thesis using n m div_half_nat[OF \m \ 0\, of n] unfolding q by(simp add: word_le_nat_alt word_div_def word_mod_def uint_nat unat_of_nat zmod_int[symmetric] zdiv_int[symmetric] word_of_nat[symmetric])(simp add: Let_def split del: if_split split: if_split_asm) qed -lemma word_test_bit_set_bits: "(BITS n. f n :: 'a :: len0 word) !! n \ n < LENGTH('a) \ f n" +lemma word_test_bit_set_bits: "(BITS n. f n :: 'a :: len word) !! n \ n < LENGTH('a) \ f n" by(auto simp add: word_set_bits_def test_bit_bl word_bl.Abs_inverse word_size) lemma word_of_int_conv_set_bits: "word_of_int i = (BITS n. i !! n)" by(rule word_eqI)(simp add: word_test_bit_set_bits) lemma word_and_mask_or_conv_and_mask: "n !! index \ (n AND mask index) OR (1 << index) = n AND mask (index + 1)" by(rule word_eqI)(auto simp add: word_ao_nth word_size nth_shiftl simp del: shiftl_1) lemma uint_and_mask_or_full: fixes n :: "'a :: len word" assumes "n !! (LENGTH('a) - 1)" and "mask1 = mask (LENGTH('a) - 1)" and "mask2 = 1 << LENGTH('a) - 1" shows "uint (n AND mask1) OR mask2 = uint n" proof - have "mask2 = uint (1 << LENGTH('a) - 1 :: 'a word)" using assms by (simp add: uint_shiftl word_size bintrunc_shiftl del: shiftl_1) hence "uint (n AND mask1) OR mask2 = uint (n AND mask1 OR (1 << LENGTH('a) - 1 :: 'a word))" by(simp add: uint_or) also have "\ = uint (n AND mask (LENGTH('a) - 1 + 1))" using assms by(simp only: word_and_mask_or_conv_and_mask) also have "\ = uint n" by simp finally show ?thesis . qed text \Division on @{typ "'a word"} is unsigned, but Scala and OCaml only have signed division and modulus.\ definition word_sdiv :: "'a :: len word \ 'a word \ 'a word" (infixl "sdiv" 70) where [code]: "x sdiv y = (let x' = sint x; y' = sint y; negative = (x' < 0) \ (y' < 0); result = abs x' div abs y' in word_of_int (if negative then -result else result))" definition word_smod :: "'a :: len word \ 'a word \ 'a word" (infixl "smod" 70) where [code]: "x smod y = (let x' = sint x; y' = sint y; negative = (x' < 0); result = abs x' mod abs y' in word_of_int (if negative then -result else result))" lemma sdiv_smod_id: "(a sdiv b) * b + (a smod b) = a" proof - note [simp] = word_sdiv_def word_smod_def have F5: "\u::'a word. - (- u) = u" by (metis minus_minus) have F7: "\v u::'a word. u + v = v + u" by(metis add.left_commute add_0_right) have F8: "\(w::'a word) (v::int) u::int. word_of_int u + word_of_int v * w = word_of_int (u + v * sint w)" by (metis word_sint.Rep_inverse wi_hom_syms(1) wi_hom_syms(3)) have "\u. u = - sint b \ word_of_int (sint a mod u + - (- u * (sint a div u))) = a" using F5 by (metis minus_minus word_sint.Rep_inverse' mult_minus_left add.commute mult_div_mod_eq [symmetric]) hence "word_of_int (sint a mod - sint b + - (sint b * (sint a div - sint b))) = a" by (metis equation_minus_iff) hence "word_of_int (sint a mod - sint b) + word_of_int (- (sint a div - sint b)) * b = a" using F8 by(metis mult.commute mult_minus_left) hence eq: "word_of_int (- (sint a div - sint b)) * b + word_of_int (sint a mod - sint b) = a" using F7 by metis show ?thesis proof(cases "sint a < 0") case True note a = this show ?thesis proof(cases "sint b < 0") case True with a show ?thesis by simp (metis F7 F8 eq minus_equation_iff minus_mult_minus mod_div_mult_eq) next case False from eq have "word_of_int (- (- sint a div sint b)) * b + word_of_int (- (- sint a mod sint b)) = a" by (metis div_minus_right mod_minus_right) with a False show ?thesis by simp qed next case False note a = this show ?thesis proof(cases "sint b < 0") case True with a eq show ?thesis by simp next case False with a show ?thesis by simp (metis wi_hom_add wi_hom_mult add.commute mult.commute word_sint.Rep_inverse add.commute mult_div_mod_eq [symmetric]) qed qed qed text \ This algorithm implements unsigned division in terms of signed division. Taken from Hacker's Delight. \ lemma divmod_via_sdivmod: fixes x y :: "'a :: len word" assumes "y \ 0" shows "(x div y, x mod y) = (if 1 << (LENGTH('a) - 1) \ y then if x < y then (0, x) else (1, x - y) else let q = ((x >> 1) sdiv y) << 1; r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" proof(cases "1 << (LENGTH('a) - 1) \ y") case True note y = this show ?thesis proof(cases "x < y") case True then have "x mod y = x" by (cases x, cases y) (simp add: word_less_def word_mod_def) thus ?thesis using True y by(simp add: word_div_lt_eq_0) next case False obtain n where n: "y = of_nat n" "n < 2 ^ LENGTH('a)" by(cases y) have "unat x < 2 ^ LENGTH('a)" by(rule unat_lt2p) also have "\ = 2 * 2 ^ (LENGTH('a) - 1)" by(metis Suc_pred len_gt_0 power_Suc One_nat_def) also have "\ \ 2 * n" using y n by(simp add: word_le_nat_alt unat_of_nat unat_p2) finally have div: "x div of_nat n = 1" using False n by(simp add: word_div_eq_1_iff not_less word_le_nat_alt unat_of_nat) moreover have "x mod y = x - x div y * y" by (metis add_diff_cancel2 word_mod_div_equality) with div n have "x mod y = x - y" by simp ultimately show ?thesis using False y n by simp qed next case False note y = this obtain n where n: "x = of_nat n" "n < 2 ^ LENGTH('a)" by(cases x) hence "int n div 2 + 2 ^ (LENGTH('a) - Suc 0) < 2 ^ LENGTH('a)" by (cases "LENGTH('a)") (simp_all, simp only: of_nat_numeral [where ?'a = int, symmetric] zdiv_int [symmetric] of_nat_power [symmetric]) with y n have "sint (x >> 1) = uint (x >> 1)" by (simp add: sint_uint sbintrunc_mod2p shiftr_div_2n) (simp add: uint_nat unat_of_nat) moreover have "uint y + 2 ^ (LENGTH('a) - Suc 0) < 2 ^ LENGTH('a)" using y by (cases "LENGTH('a)") (simp_all add: not_le word_2p_lem word_size) then have "sint y = uint y" by (simp add: sint_uint sbintrunc_mod2p) ultimately show ?thesis using y by(subst div_half_word[OF assms])(simp add: word_sdiv_def uint_div[symmetric]) qed text \More implementations tailored towards target-language implementations\ context includes integer.lifting begin -lift_definition word_of_integer :: "integer \ 'a :: len0 word" is word_of_int . +lift_definition word_of_integer :: "integer \ 'a :: len word" is word_of_int . lemma word_of_integer_code [code]: "word_of_integer n = word_of_int (int_of_integer n)" by(simp add: word_of_integer.rep_eq) end lemma word_of_int_code [code abstract]: - "uint (word_of_int x :: 'a word) = x AND bin_mask (LENGTH('a :: len0))" + "uint (word_of_int x :: 'a word) = x AND bin_mask (LENGTH('a :: len))" by(simp add: uint_word_of_int and_bin_mask_conv_mod) context fixes f :: "nat \ bool" begin -fun set_bits_aux :: "'a word \ nat \ 'a :: len0 word" +fun set_bits_aux :: "'a word \ nat \ 'a :: len word" where "set_bits_aux w 0 = w" | "set_bits_aux w (Suc n) = set_bits_aux ((w << 1) OR (if f n then 1 else 0)) n" lemma set_bits_aux_conv: "set_bits_aux w n = (w << n) OR (set_bits f AND mask n)" apply(induct w n rule: set_bits_aux.induct) apply(auto 4 3 intro: word_eqI simp add: word_ao_nth nth_shiftl test_bit.eq_norm word_size not_less less_Suc_eq) done corollary set_bits_conv_set_bits_aux: "set_bits f = (set_bits_aux 0 (LENGTH('a)) :: 'a :: len word)" by(simp add: set_bits_aux_conv) end lemma word_of_int_via_signed: fixes mask assumes mask_def: "mask = bin_mask (LENGTH('a))" and shift_def: "shift = 1 << LENGTH('a)" and index_def: "index = LENGTH('a) - 1" and overflow_def:"overflow = 1 << (LENGTH('a) - 1)" and least_def: "least = - overflow" shows "(word_of_int i :: 'a :: len word) = (let i' = i AND mask in if i' !! index then if i' - shift < least \ overflow \ i' - shift then arbitrary1 i' else word_of_int (i' - shift) else if i' < least \ overflow \ i' then arbitrary2 i' else word_of_int i')" proof - define i' where "i' = i AND mask" have "shift = mask + 1" unfolding assms by(simp add: bin_mask_p1_conv_shift) hence "i' < shift" by(simp add: mask_def i'_def int_and_le) show ?thesis proof(cases "i' !! index") case True hence unf: "i' = overflow OR i'" unfolding assms i'_def by(auto intro!: bin_eqI simp add: bin_nth_ops) have "overflow \ i'" by(subst unf)(rule le_int_or, simp add: bin_sign_and assms i'_def) hence "i' - shift < least \ False" unfolding assms by(cases "LENGTH('a)")(simp_all add: not_less Bit_def) moreover have "overflow \ i' - shift \ False" using \i' < shift\ unfolding assms by(cases "LENGTH('a)")(auto simp add: not_le Bit_def elim: less_le_trans) moreover have "word_of_int (i' - shift) = (word_of_int i :: 'a word)" using \i' < shift\ by(auto intro!: word_eqI simp add: i'_def shift_def mask_def bin_nth_ops bin_nth_minus_p2 bin_sign_and) ultimately show ?thesis using True by(simp add: Let_def i'_def) next case False hence "i' = i AND bin_mask (LENGTH('a) - 1)" unfolding assms i'_def by(clarsimp simp add: i'_def bin_nth_ops intro!: bin_eqI)(cases "LENGTH('a)", auto simp add: less_Suc_eq) also have "\ \ bin_mask (LENGTH('a) - 1)" by(rule int_and_le) simp also have "\ < overflow" unfolding overflow_def by(simp add: bin_mask_p1_conv_shift[symmetric]) also have "least \ 0" unfolding least_def overflow_def by simp have "0 \ i'" by(simp add: i'_def mask_def bin_mask_ge0) hence "least \ i'" using \least \ 0\ by simp moreover have "word_of_int i' = (word_of_int i :: 'a word)" by(rule word_eqI)(auto simp add: i'_def bin_nth_ops mask_def) ultimately show ?thesis using False by(simp add: Let_def i'_def) qed qed text \Quickcheck conversion functions\ notation scomp (infixl "\\" 60) definition qc_random_cnv :: "(natural \ 'a::term_of) \ natural \ Random.seed \ ('a \ (unit \ Code_Evaluation.term)) \ Random.seed" where "qc_random_cnv a_of_natural i = Random.range (i + 1) \\ (\k. Pair ( let n = a_of_natural k in (n, \_. Code_Evaluation.term_of n)))" no_notation scomp (infixl "\\" 60) definition qc_exhaustive_cnv :: "(natural \ 'a) \ ('a \ (bool \ term list) option) \ natural \ (bool \ term list) option" where "qc_exhaustive_cnv a_of_natural f d = Quickcheck_Exhaustive.exhaustive (%x. f (a_of_natural x)) d" definition qc_full_exhaustive_cnv :: "(natural \ ('a::term_of)) \ ('a \ (unit \ term) \ (bool \ term list) option) \ natural \ (bool \ term list) option" where "qc_full_exhaustive_cnv a_of_natural f d = Quickcheck_Exhaustive.full_exhaustive (%(x, xt). f (a_of_natural x, %_. Code_Evaluation.term_of (a_of_natural x))) d" declare [[quickcheck_narrowing_ghc_options = "-XTypeSynonymInstances"]] definition qc_narrowing_drawn_from :: "'a list \ integer \ _" where "qc_narrowing_drawn_from xs = foldr Quickcheck_Narrowing.sum (map Quickcheck_Narrowing.cons (butlast xs)) (Quickcheck_Narrowing.cons (last xs))" locale quickcheck_narrowing_samples = fixes a_of_integer :: "integer \ 'a \ 'a :: {partial_term_of, term_of}" and zero :: "'a" and tr :: "typerep" begin function narrowing_samples :: "integer \ 'a list" where "narrowing_samples i = (if i > 0 then let (a, a') = a_of_integer i in narrowing_samples (i - 1) @ [a, a'] else [zero])" by pat_completeness auto termination including integer.lifting proof(relation "measure nat_of_integer") fix i :: integer assume "0 < i" thus "(i - 1, i) \ measure nat_of_integer" by simp(transfer, simp) qed simp definition partial_term_of_sample :: "integer \ 'a" where "partial_term_of_sample i = (if i < 0 then undefined else if i = 0 then zero else if i mod 2 = 0 then snd (a_of_integer (i div 2)) else fst (a_of_integer (i div 2 + 1)))" lemma partial_term_of_code: "partial_term_of (ty :: 'a itself) (Quickcheck_Narrowing.Narrowing_variable p t) \ Code_Evaluation.Free (STR ''_'') tr" "partial_term_of (ty :: 'a itself) (Quickcheck_Narrowing.Narrowing_constructor i []) \ Code_Evaluation.term_of (partial_term_of_sample i)" by (rule partial_term_of_anything)+ end lemmas [code] = quickcheck_narrowing_samples.narrowing_samples.simps quickcheck_narrowing_samples.partial_term_of_sample_def text \ The separate code target \SML_word\ collects setups for the code generator that PolyML does not provide. \ setup \Code_Target.add_derived_target ("SML_word", [(Code_ML.target_SML, I)])\ code_identifier code_module Code_Target_Word_Base \ (SML) Word and (Haskell) Word and (OCaml) Word and (Scala) Word export_code sbintrunc bin_mask in SML end diff --git a/thys/Refine_Monadic/examples/WordRefine.thy b/thys/Refine_Monadic/examples/WordRefine.thy --- a/thys/Refine_Monadic/examples/WordRefine.thy +++ b/thys/Refine_Monadic/examples/WordRefine.thy @@ -1,80 +1,79 @@ section \Machine Words\ theory WordRefine imports "../Refine_Monadic" "HOL-Word.Word" begin text \This theory provides a simple example to show refinement of natural numbers to machine words. The setup is not yet very elaborated, but shows the direction to go. \ subsection \Setup\ definition [simp]: "word_nat_rel \ build_rel (unat) (\_. True)" lemma word_nat_RELEATES[refine_dref_RELATES]: "RELATES word_nat_rel" by (simp add: RELATES_def) lemma [simp, relator_props]: "single_valued word_nat_rel" unfolding word_nat_rel_def by blast lemma [simp]: "single_valuedp (\c a. a = unat c)" by (rule single_valuedpI) blast lemma [simp, relator_props]: "single_valued (converse word_nat_rel)" - apply (auto intro!: injI) - by (metis order_antisym order_eq_refl word_le_nat_alt) + by (auto intro: injI) lemmas [refine_hsimp] = word_less_nat_alt word_le_nat_alt unat_sub iffD1[OF unat_add_lem] subsection \Example\ type_synonym word32 = "32 word" definition test :: "nat \ nat \ nat set nres" where "test x0 y0 \ do { let S={}; (S,_,_) \ WHILE (\(S,x,y). x>0) (\(S,x,y). do { let S=S\{y}; let x=x - 1; ASSERT (y0 \ test x0 y0 \ SPEC (\S. S={y0 .. y0 + x0 - 1})" \ \Choosen pre-condition to get least trouble when proving\ unfolding test_def apply (intro WHILE_rule[where I="\(S,x,y). x+y=x0+y0 \ x\x0 \ S={y0 .. y0 + (x0-x) - 1}"] refine_vcg) by auto definition test_impl :: "word32 \ word32 \ word32 set nres" where "test_impl x y \ do { let S={}; (S,_,_) \ WHILE (\(S,x,y). x>0) (\(S,x,y). do { let S=S\{y}; let x=x - 1; let y=y + 1; RETURN (S,x,y) }) (S,x,y); RETURN S }" lemma test_impl_refine: assumes "x'+y'<2^LENGTH(32)" assumes "(x,x')\word_nat_rel" assumes "(y,y')\word_nat_rel" shows "test_impl x y \ \(\word_nat_rel\set_rel) (test x' y')" proof - from assms show ?thesis unfolding test_impl_def test_def apply (refine_rcg) apply (refine_dref_type) apply (auto simp: refine_hsimp refine_rel_defs) done qed end diff --git a/thys/SPARCv8/SparcModel_MMU/Sparc_Execution.thy b/thys/SPARCv8/SparcModel_MMU/Sparc_Execution.thy --- a/thys/SPARCv8/SparcModel_MMU/Sparc_Execution.thy +++ b/thys/SPARCv8/SparcModel_MMU/Sparc_Execution.thy @@ -1,431 +1,431 @@ (* * Copyright 2016, NTU * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * Author: Zhe Hou, David Sanan. *) theory Sparc_Execution imports Main Sparc_Instruction Sparc_State Sparc_Types "HOL-Eisbach.Eisbach_Tools" begin primrec sum :: "nat \ nat" where "sum 0 = 0" | "sum (Suc n) = Suc n + sum n" definition select_trap :: "unit \ ('a,unit) sparc_state_monad" where "select_trap _ \ do traps \ gets (\s. (get_trap_set s)); rt_val \ gets (\s. (reset_trap_val s)); psr_val \ gets (\s. (cpu_reg_val PSR s)); et_val \ gets (\s. (get_ET psr_val)); modify (\s. (emp_trap_set s)); if rt_val = True then \ \ignore \ET\, and leave \tt\ unchaged\ return () else if et_val = 0 then \ \go to error mode, machine needs reset\ do set_err_mode True; set_exe_mode False; fail () od \ \By the SPARCv8 manual only 1 of the following traps could be in traps.\ else if data_store_error \ traps then do write_cpu_tt (0b00101011::word8); return () od else if instruction_access_error \ traps then do write_cpu_tt (0b00100001::word8); return () od else if r_register_access_error \ traps then do write_cpu_tt (0b00100000::word8); return () od else if instruction_access_exception \ traps then do write_cpu_tt (0b00000001::word8); return () od else if privileged_instruction \ traps then do write_cpu_tt (0b00000011::word8); return () od else if illegal_instruction \ traps then do write_cpu_tt (0b00000010::word8); return () od else if fp_disabled \ traps then do write_cpu_tt (0b00000100::word8); return () od else if cp_disabled \ traps then do write_cpu_tt (0b00100100::word8); return () od else if unimplemented_FLUSH \ traps then do write_cpu_tt (0b00100101::word8); return () od else if window_overflow \ traps then do write_cpu_tt (0b00000101::word8); return () od else if window_underflow \ traps then do write_cpu_tt (0b00000110::word8); return () od else if mem_address_not_aligned \ traps then do write_cpu_tt (0b00000111::word8); return () od else if fp_exception \ traps then do write_cpu_tt (0b00001000::word8); return () od else if cp_exception \ traps then do write_cpu_tt (0b00101000::word8); return () od else if data_access_error \ traps then do write_cpu_tt (0b00101001::word8); return () od else if data_access_exception \ traps then do write_cpu_tt (0b00001001::word8); return () od else if tag_overflow \ traps then do write_cpu_tt (0b00001010::word8); return () od else if division_by_zero \ traps then do write_cpu_tt (0b00101010::word8); return () od else if trap_instruction \ traps then do ticc_trap_type \ gets (\s. (ticc_trap_type_val s)); write_cpu_tt (word_cat (1::word1) ticc_trap_type); return () od \<^cancel>\else if interrupt_level > 0 then\ \ \We don't consider \interrupt_level\\ else return () od" -definition exe_trap_st_pc :: "unit \ ('a::len0,unit) sparc_state_monad" +definition exe_trap_st_pc :: "unit \ ('a::len,unit) sparc_state_monad" where "exe_trap_st_pc _ \ do annul \ gets (\s. (annul_val s)); pc_val \ gets (\s. (cpu_reg_val PC s)); npc_val \ gets (\s. (cpu_reg_val nPC s)); curr_win \ get_curr_win(); if annul = False then do write_reg pc_val curr_win (word_of_int 17); write_reg npc_val curr_win (word_of_int 18); return () od else \ \\annul = True\\ do write_reg npc_val curr_win (word_of_int 17); write_reg (npc_val + 4) curr_win (word_of_int 18); set_annul False; return () od od" -definition exe_trap_wr_pc :: "unit \ ('a::len0,unit) sparc_state_monad" +definition exe_trap_wr_pc :: "unit \ ('a::len,unit) sparc_state_monad" where "exe_trap_wr_pc _ \ do psr_val \ gets (\s. (cpu_reg_val PSR s)); new_psr_val \ gets (\s. (update_S (1::word1) psr_val)); write_cpu new_psr_val PSR; reset_trap \ gets (\s. (reset_trap_val s)); tbr_val \ gets (\s. (cpu_reg_val TBR s)); if reset_trap = False then do write_cpu tbr_val PC; write_cpu (tbr_val + 4) nPC; return () od else \ \\reset_trap = True\\ do write_cpu 0 PC; write_cpu 4 nPC; set_reset_trap False; return () od od" -definition execute_trap :: "unit \ ('a::len0,unit) sparc_state_monad" +definition execute_trap :: "unit \ ('a::len,unit) sparc_state_monad" where "execute_trap _ \ do select_trap(); err_mode \ gets (\s. (err_mode_val s)); if err_mode = True then \ \The SparcV8 manual doesn't say what to do.\ return () else do psr_val \ gets (\s. (cpu_reg_val PSR s)); s_val \ gets (\s. ((ucast (get_S psr_val))::word1)); curr_win \ get_curr_win(); new_cwp \ gets (\s. ((word_of_int (((uint curr_win) - 1) mod NWINDOWS)))::word5); new_psr_val \ gets (\s. (update_PSR_exe_trap new_cwp (0::word1) s_val psr_val)); write_cpu new_psr_val PSR; exe_trap_st_pc(); exe_trap_wr_pc(); return () od od" -definition dispatch_instruction :: "instruction \ ('a::len0,unit) sparc_state_monad" +definition dispatch_instruction :: "instruction \ ('a::len,unit) sparc_state_monad" where "dispatch_instruction instr \ let instr_name = fst instr in do traps \ gets (\s. (get_trap_set s)); if traps = {} then if instr_name \ {load_store_type LDSB,load_store_type LDUB, load_store_type LDUBA,load_store_type LDUH,load_store_type LD, load_store_type LDA,load_store_type LDD} then load_instr instr else if instr_name \ {load_store_type STB,load_store_type STH, load_store_type ST,load_store_type STA,load_store_type STD} then store_instr instr else if instr_name \ {sethi_type SETHI} then sethi_instr instr else if instr_name \ {nop_type NOP} then nop_instr instr else if instr_name \ {logic_type ANDs,logic_type ANDcc,logic_type ANDN, logic_type ANDNcc,logic_type ORs,logic_type ORcc,logic_type ORN, logic_type XORs,logic_type XNOR} then logical_instr instr else if instr_name \ {shift_type SLL,shift_type SRL,shift_type SRA} then shift_instr instr else if instr_name \ {arith_type ADD,arith_type ADDcc,arith_type ADDX} then add_instr instr else if instr_name \ {arith_type SUB,arith_type SUBcc,arith_type SUBX} then sub_instr instr else if instr_name \ {arith_type UMUL,arith_type SMUL,arith_type SMULcc} then mul_instr instr else if instr_name \ {arith_type UDIV,arith_type UDIVcc,arith_type SDIV} then div_instr instr else if instr_name \ {ctrl_type SAVE,ctrl_type RESTORE} then save_restore_instr instr else if instr_name \ {call_type CALL} then call_instr instr else if instr_name \ {ctrl_type JMPL} then jmpl_instr instr else if instr_name \ {ctrl_type RETT} then rett_instr instr else if instr_name \ {sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM, sreg_type RDTBR} then read_state_reg_instr instr else if instr_name \ {sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM, sreg_type WRTBR} then write_state_reg_instr instr else if instr_name \ {load_store_type FLUSH} then flush_instr instr else if instr_name \ {bicc_type BE,bicc_type BNE,bicc_type BGU, bicc_type BLE,bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA,bicc_type BN} then branch_instr instr else fail () else return () od" definition supported_instruction :: "sparc_operation \ bool" where "supported_instruction instr \ if instr \ {load_store_type LDSB,load_store_type LDUB,load_store_type LDUBA, load_store_type LDUH,load_store_type LD,load_store_type LDA, load_store_type LDD, load_store_type STB,load_store_type STH,load_store_type ST, load_store_type STA,load_store_type STD, sethi_type SETHI, nop_type NOP, logic_type ANDs,logic_type ANDcc,logic_type ANDN,logic_type ANDNcc, logic_type ORs,logic_type ORcc,logic_type ORN,logic_type XORs, logic_type XNOR, shift_type SLL,shift_type SRL,shift_type SRA, arith_type ADD,arith_type ADDcc,arith_type ADDX, arith_type SUB,arith_type SUBcc,arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, ctrl_type RETT, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN} then True else False " -definition execute_instr_sub1 :: "instruction \ ('a::len0,unit) sparc_state_monad" +definition execute_instr_sub1 :: "instruction \ ('a::len,unit) sparc_state_monad" where "execute_instr_sub1 instr \ do instr_name \ gets (\s. (fst instr)); traps2 \ gets (\s. (get_trap_set s)); if traps2 = {} \ instr_name \ {call_type CALL,ctrl_type RETT,ctrl_type JMPL, bicc_type BE,bicc_type BNE,bicc_type BGU, bicc_type BLE,bicc_type BL,bicc_type BGE, bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC, bicc_type BA,bicc_type BN} then do npc_val \ gets (\s. (cpu_reg_val nPC s)); write_cpu npc_val PC; write_cpu (npc_val + 4) nPC; return () od else return () od" -definition execute_instruction :: "unit \ ('a::len0,unit) sparc_state_monad" +definition execute_instruction :: "unit \ ('a::len,unit) sparc_state_monad" where "execute_instruction _ \ do traps \ gets (\s. (get_trap_set s)); if traps = {} then do exe_mode \ gets (\s. (exe_mode_val s)); if exe_mode = True then do modify (\s. (delayed_pool_write s)); fetch_result \ gets (\s. (fetch_instruction s)); case fetch_result of Inl e1 \ (do \ \Memory address in PC is not aligned.\ \ \Actually, SparcV8 manual doens't check alignment here.\ raise_trap instruction_access_exception; return () od) | Inr v1 \ (do dec \ gets (\s. (decode_instruction v1)); case dec of Inl e2 \ (\ \Instruction is ill-formatted.\ fail () ) | Inr v2 \ (do instr \ gets (\s. (v2)); annul \ gets (\s. (annul_val s)); if annul = False then do dispatch_instruction instr; execute_instr_sub1 instr; return () od else \ \\annul \ False\\ do set_annul False; npc_val \ gets (\s. (cpu_reg_val nPC s)); write_cpu npc_val PC; write_cpu (npc_val + 4) nPC; return () od od) od) od else return () \ \Not in \execute_mode\.\ od else \ \traps is not empty, which means \trap = 1\.\ do execute_trap(); return () od od" -definition NEXT :: "('a::len0)sparc_state \ ('a)sparc_state option" +definition NEXT :: "('a::len)sparc_state \ ('a)sparc_state option" where "NEXT s \ case execute_instruction () s of (_,True) \ None | (s',False) \ Some (snd s')" -definition good_context :: "('a::len0) sparc_state \ bool" +definition good_context :: "('a::len) sparc_state \ bool" where "good_context s \ let traps = get_trap_set s; psr_val = cpu_reg_val PSR s; et_val = get_ET psr_val; rt_val = reset_trap_val s in if traps \ {} \ rt_val = False \ et_val = 0 then False \ \enter \error_mode\ in \select_traps\.\ else let s' = delayed_pool_write s in case fetch_instruction s' of \ \\instruction_access_exception\ is handled in the next state.\ Inl _ \ True |Inr v \ ( case decode_instruction v of Inl _ \ False |Inr instr \ ( let annul = annul_val s' in if annul = True then True else \ \\annul = False\\ if supported_instruction (fst instr) then \ \The only instruction that could fail is \RETT\.\ if (fst instr) = ctrl_type RETT then let curr_win_r = (get_CWP (cpu_reg_val PSR s')); new_cwp_int_r = (((uint curr_win_r) + 1) mod NWINDOWS); wim_val_r = cpu_reg_val WIM s'; psr_val_r = cpu_reg_val PSR s'; et_val_r = get_ET psr_val_r; s_val_r = (ucast (get_S psr_val_r))::word1; op_list_r = snd instr; addr_r = get_addr (snd instr) s' in if et_val_r = 1 then True else if s_val_r = 0 then False else if (get_WIM_bit (nat new_cwp_int_r) wim_val_r) \ 0 then False else if ((AND) addr_r (0b00000000000000000000000000000011::word32)) \ 0 then False else True else True else False \ \Unsupported instruction.\ ) ) " -function (sequential) seq_exec:: "nat \ ('a::len0,unit) sparc_state_monad" +function (sequential) seq_exec:: "nat \ ('a::len,unit) sparc_state_monad" where "seq_exec 0 = return ()" | "seq_exec n = (do execute_instruction(); (seq_exec (n-1)) od) " by pat_completeness auto termination by lexicographic_order type_synonym leon3_state = "(word_length5) sparc_state" type_synonym ('e) leon3_state_monad = "(leon3_state, 'e) det_monad" definition execute_leon3_instruction:: "unit \ (unit) leon3_state_monad" where "execute_leon3_instruction \ execute_instruction" definition seq_exec_leon3:: "nat \ (unit) leon3_state_monad" where "seq_exec_leon3 \ seq_exec" end diff --git a/thys/SPARCv8/SparcModel_MMU/Sparc_Instruction.thy b/thys/SPARCv8/SparcModel_MMU/Sparc_Instruction.thy --- a/thys/SPARCv8/SparcModel_MMU/Sparc_Instruction.thy +++ b/thys/SPARCv8/SparcModel_MMU/Sparc_Instruction.thy @@ -1,2788 +1,2788 @@ (* * Copyright 2016, NTU * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * Author: Zhe Hou, David Sanan. *) section \SPARC instruction model\ theory Sparc_Instruction imports Main Sparc_Types Sparc_State "HOL-Eisbach.Eisbach_Tools" begin text\ This theory provides a formal model for assembly instruction to be executed in the model. An instruction is defined as a tuple composed of a @{term sparc_operation} element, defining the operation the instruction carries out, and a list of operands @{term inst_operand}. @{term inst_operand} can be a user register @{term user_reg} or a memory address @{term mem_add_type}. \ datatype inst_operand = W5 word5 |W30 word30 |W22 word22 |Cond word4 |Flag word1 |Asi asi_type |Simm13 word13 |Opf word9 |Imm7 word7 primrec get_operand_w5::"inst_operand \ word5" where "get_operand_w5 (W5 r) = r" primrec get_operand_w30::"inst_operand \ word30" where "get_operand_w30 (W30 r) = r" primrec get_operand_w22::"inst_operand \ word22" where "get_operand_w22 (W22 r) = r" primrec get_operand_cond::"inst_operand \ word4" where "get_operand_cond (Cond r) = r" primrec get_operand_flag::"inst_operand \ word1" where "get_operand_flag (Flag r) = r" primrec get_operand_asi::"inst_operand \ asi_type" where "get_operand_asi (Asi r) = r" primrec get_operand_simm13::"inst_operand \ word13" where "get_operand_simm13 (Simm13 r) = r" primrec get_operand_opf::"inst_operand \ word9" where "get_operand_opf (Opf r) = r" primrec get_operand_imm7:: "inst_operand \ word7" where "get_operand_imm7 (Imm7 r) = r" type_synonym instruction = "(sparc_operation \ inst_operand list)" definition get_op::"word32 \ int" where "get_op w \ uint (w >> 30)" definition get_op2::"word32 \ int" where "get_op2 w \ let mask_op2 = 0b00000001110000000000000000000000 in uint (((AND) mask_op2 w) >> 22)" definition get_op3::"word32 \ int" where "get_op3 w \ let mask_op3 = 0b00000001111110000000000000000000 in uint (((AND) mask_op3 w) >> 19)" definition get_disp30::"word32 \ int" where "get_disp30 w \ let mask_disp30 = 0b00111111111111111111111111111111 in uint ((AND) mask_disp30 w)" definition get_a::"word32 \ int" where "get_a w \ let mask_a = 0b00100000000000000000000000000000 in uint (((AND) mask_a w) >> 29)" definition get_cond::"word32 \ int" where "get_cond w \ let mask_cond = 0b00011110000000000000000000000000 in uint (((AND) mask_cond w) >> 25)" definition get_disp_imm22::"word32 \ int" where "get_disp_imm22 w \ let mask_disp_imm22 = 0b00000000001111111111111111111111 in uint ((AND) mask_disp_imm22 w)" definition get_rd::"word32 \ int" where "get_rd w \ let mask_rd = 0b00111110000000000000000000000000 in uint (((AND) mask_rd w) >> 25)" definition get_rs1::"word32 \ int" where "get_rs1 w \ let mask_rs1 = 0b00000000000001111100000000000000 in uint (((AND) mask_rs1 w) >> 14)" definition get_i::"word32 \ int" where "get_i w \ let mask_i = 0b00000000000000000010000000000000 in uint (((AND) mask_i w) >> 13)" definition get_opf::"word32 \ int" where "get_opf w \ let mask_opf = 0b00000000000000000011111111100000 in uint (((AND) mask_opf w) >> 5)" definition get_rs2::"word32 \ int" where "get_rs2 w \ let mask_rs2 = 0b00000000000000000000000000011111 in uint ((AND) mask_rs2 w)" definition get_simm13::"word32 \ int" where "get_simm13 w \ let mask_simm13 = 0b00000000000000000001111111111111 in uint ((AND) mask_simm13 w)" definition get_asi::"word32 \ int" where "get_asi w \ let mask_asi = 0b00000000000000000001111111100000 in uint (((AND) mask_asi w) >> 5)" definition get_trap_cond:: "word32 \ int" where "get_trap_cond w \ let mask_cond = 0b00011110000000000000000000000000 in uint (((AND) mask_cond w) >> 25)" definition get_trap_imm7:: "word32 \ int" where "get_trap_imm7 w \ let mask_imm7 = 0b00000000000000000000000001111111 in uint ((AND) mask_imm7 w)" definition parse_instr_f1::"word32 \ (Exception list + instruction)" where \ \\CALL\, with a single operand \disp30+"00"\\ "parse_instr_f1 w \ Inr (call_type CALL,[W30 (word_of_int (get_disp30 w))])" definition parse_instr_f2::"word32 \ (Exception list + instruction)" where "parse_instr_f2 w \ let op2 = get_op2 w in if op2 = uint(0b100::word3) then \ \\SETHI\ or \NOP\\ let rd = get_rd w in let imm22 = get_disp_imm22 w in if rd = 0 \ imm22 = 0 then \ \\NOP\\ Inr (nop_type NOP,[]) else \ \\SETHI\, with operands \[imm22,rd]\\ Inr (sethi_type SETHI,[(W22 (word_of_int imm22)), (W5 (word_of_int rd))]) else if op2 = uint(0b010::word3) then \ \\Bicc\, with operands \[a,disp22]\\ let cond = get_cond w in let flaga = Flag (word_of_int (get_a w)) in let disp22 = W22 (word_of_int (get_disp_imm22 w)) in if cond = uint(0b0001::word4) then \ \\BE\\ Inr (bicc_type BE,[flaga,disp22]) else if cond = uint(0b1001::word4) then \ \\BNE\\ Inr (bicc_type BNE,[flaga,disp22]) else if cond = uint(0b1100::word4) then \ \\BGU\\ Inr (bicc_type BGU,[flaga,disp22]) else if cond = uint(0b0010::word4) then \ \\BLE\\ Inr (bicc_type BLE,[flaga,disp22]) else if cond = uint(0b0011::word4) then \ \\BL\\ Inr (bicc_type BL,[flaga,disp22]) else if cond = uint(0b1011::word4) then \ \\BGE\\ Inr (bicc_type BGE,[flaga,disp22]) else if cond = uint(0b0110::word4) then \ \\BNEG\\ Inr (bicc_type BNEG,[flaga,disp22]) else if cond = uint(0b1010::word4) then \ \\BG\\ Inr (bicc_type BG,[flaga,disp22]) else if cond = uint(0b0101::word4) then \ \\BCS\\ Inr (bicc_type BCS,[flaga,disp22]) else if cond = uint(0b0100::word4) then \ \\BLEU\\ Inr (bicc_type BLEU,[flaga,disp22]) else if cond = uint(0b1101::word4) then \ \\BCC\\ Inr (bicc_type BCC,[flaga,disp22]) else if cond = uint(0b1000::word4) then \ \\BA\\ Inr (bicc_type BA,[flaga,disp22]) else if cond = uint(0b0000::word4) then \ \\BN\\ Inr (bicc_type BN,[flaga,disp22]) else if cond = uint(0b1110::word4) then \ \\BPOS\\ Inr (bicc_type BPOS,[flaga,disp22]) else if cond = uint(0b1111::word4) then \ \\BVC\\ Inr (bicc_type BVC,[flaga,disp22]) else if cond = uint(0b0111::word4) then \ \\BVS\\ Inr (bicc_type BVS,[flaga,disp22]) else Inl [invalid_cond_f2] else Inl [invalid_op2_f2] " text \We don't consider floating-point operations, so we don't consider the third type of format 3.\ definition parse_instr_f3::"word32 \ (Exception list + instruction)" where "parse_instr_f3 w \ let this_op = get_op w in let rd = get_rd w in let op3 = get_op3 w in let rs1 = get_rs1 w in let flagi = get_i w in let asi = get_asi w in let rs2 = get_rs2 w in let simm13 = get_simm13 w in if this_op = uint(0b11::word2) then \ \Load and Store\ \ \If an instruction accesses alternative space but \flagi = 1\,\ \ \may need to throw a trap.\ if op3 = uint(0b001001::word6) then \ \\LDSB\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type LDSB,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type LDSB,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b011001::word6) then \ \\LDSBA\\ Inr (load_store_type LDSBA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else if op3 = uint(0b001010::word6) then \ \\LDSH\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type LDSH,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type LDSH,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b011010::word6) then \ \\LDSHA\\ Inr (load_store_type LDSHA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else if op3 = uint(0b000001::word6) then \ \\LDUB\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type LDUB,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type LDUB,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b010001::word6) then \ \\LDUBA\\ Inr (load_store_type LDUBA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else if op3 = uint(0b000010::word6) then \ \\LDUH\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type LDUH,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type LDUH,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b010010::word6) then \ \\LDUHA\\ Inr (load_store_type LDUHA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else if op3 = uint(0b000000::word6) then \ \\LD\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type LD,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type LD,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b010000::word6) then \ \\LDA\\ Inr (load_store_type LDA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else if op3 = uint(0b000011::word6) then \ \\LDD\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type LDD,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type LDD,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b010011::word6) then \ \\LDDA\\ Inr (load_store_type LDDA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else if op3 = uint(0b001101::word6) then \ \\LDSTUB\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type LDSTUB,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type LDSTUB,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b011101::word6) then \ \\LDSTUBA\\ Inr (load_store_type LDSTUBA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else if op3 = uint(0b000101::word6) then \ \\STB\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type STB,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type STB,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b010101::word6) then \ \\STBA\\ Inr (load_store_type STBA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else if op3 = uint(0b000110::word6) then \ \\STH\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type STH,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type STH,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b010110::word6) then \ \\STHA\\ Inr (load_store_type STHA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else if op3 = uint(0b000100::word6) then \ \\ST\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type ST,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type ST,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b010100::word6) then \ \\STA\\ Inr (load_store_type STA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else if op3 = uint(0b000111::word6) then \ \\STD\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type STD,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type STD,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b010111::word6) then \ \\STDA\\ Inr (load_store_type STDA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else if op3 = uint(0b001111::word6) then \ \\SWAP\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type SWAP,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type SWAP,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b011111::word6) then \ \\SWAPA\\ Inr (load_store_type SWAPA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else Inl [invalid_op3_f3_op11] else if this_op = uint(0b10::word2) then \ \Others\ if op3 = uint(0b111000::word6) then \ \\JMPL\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (ctrl_type JMPL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (ctrl_type JMPL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b111001::word6) then \ \\RETT\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ctrl_type RETT,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,simm13]\\ Inr (ctrl_type RETT,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13))]) \ \The following are Read and Write instructions,\ \ \only return \[rs1,rd]\ as operand.\ else if op3 = uint(0b101000::word6) \ rs1 \ 0 then \ \\RDASR\\ if rs1 = uint(0b01111::word6) \ rd = 0 then \ \\STBAR\ is a special case of \RDASR\\ Inr (load_store_type STBAR,[]) else Inr (sreg_type RDASR,[(W5 (word_of_int rs1)), (W5 (word_of_int rd))]) else if op3 = uint(0b101000::word6) \ rs1 = 0 then \ \\RDY\\ Inr (sreg_type RDY,[(W5 (word_of_int rs1)), (W5 (word_of_int rd))]) else if op3 = uint(0b101001::word6) then \ \\RDPSR\\ Inr (sreg_type RDPSR,[(W5 (word_of_int rs1)), (W5 (word_of_int rd))]) else if op3 = uint(0b101010::word6) then \ \\RDWIM\\ Inr (sreg_type RDWIM,[(W5 (word_of_int rs1)), (W5 (word_of_int rd))]) else if op3 = uint(0b101011::word6) then \ \\RDTBR\\ Inr (sreg_type RDTBR,[(W5 (word_of_int rs1)), (W5 (word_of_int rd))]) else if op3 = uint(0b110000::word6) \ rd \ 0 then \ \\WRASR\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (sreg_type WRASR,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (sreg_type WRASR,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b110000::word6) \ rd = 0 then \ \\WRY\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (sreg_type WRY,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (sreg_type WRY,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b110001::word6) then \ \\WRPSR\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (sreg_type WRPSR,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (sreg_type WRPSR,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b110010::word6) then \ \\WRWIM\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (sreg_type WRWIM,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (sreg_type WRWIM,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b110011::word6) then \ \\WRTBR\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (sreg_type WRTBR,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (sreg_type WRTBR,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) \ \\FLUSH\ instruction\ else if op3 = uint(0b111011::word6) then \ \\FLUSH\\ if flagi = 0 then \ \return \[1,rs1,rs2]\\ Inr (load_store_type FLUSH,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,simm13]\\ Inr (load_store_type FLUSH,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13))]) \ \The following are arithmetic instructions.\ else if op3 = uint(0b000001::word6) then \ \\AND\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type ANDs,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type ANDs,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b010001::word6) then \ \\ANDcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type ANDcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type ANDcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b000101::word6) then \ \\ANDN\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type ANDN,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type ANDN,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b010101::word6) then \ \\ANDNcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type ANDNcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type ANDNcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b000010::word6) then \ \\OR\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type ORs,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type ORs,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b010010::word6) then \ \\ORcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type ORcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type ORcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b000110::word6) then \ \\ORN\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type ORN,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type ORN,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b010110::word6) then \ \\ORNcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type ORNcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type ORNcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b000011::word6) then \ \\XORs\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type XORs,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type XORs,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b010011::word6) then \ \\XORcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type XORcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type XORcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b000111::word6) then \ \\XNOR\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type XNOR,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type XNOR,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b010111::word6) then \ \\XNORcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type XNORcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type XNORcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b100101::word6) then \ \\SLL\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (shift_type SLL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,shcnt,rd]\\ let shcnt = rs2 in Inr (shift_type SLL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int shcnt)), (W5 (word_of_int rd))]) else if op3 = uint (0b100110::word6) then \ \\SRL\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (shift_type SRL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,shcnt,rd]\\ let shcnt = rs2 in Inr (shift_type SRL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int shcnt)), (W5 (word_of_int rd))]) else if op3 = uint(0b100111::word6) then \ \\SRA\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (shift_type SRA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,shcnt,rd]\\ let shcnt = rs2 in Inr (shift_type SRA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int shcnt)), (W5 (word_of_int rd))]) else if op3 = uint(0b000000::word6) then \ \\ADD\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type ADD,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type ADD,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b010000::word6) then \ \\ADDcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type ADDcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type ADDcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b001000::word6) then \ \\ADDX\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type ADDX,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type ADDX,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b011000::word6) then \ \\ADDXcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type ADDXcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type ADDXcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b100000::word6) then \ \\TADDcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type TADDcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type TADDcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b100010::word6) then \ \\TADDccTV\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type TADDccTV,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type TADDccTV,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b000100::word6) then \ \\SUB\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type SUB,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type SUB,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b010100::word6) then \ \\SUBcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type SUBcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type SUBcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b001100::word6) then \ \\SUBX\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type SUBX,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type SUBX,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b011100::word6) then \ \\SUBXcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type SUBXcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type SUBXcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b100001::word6) then \ \\TSUBcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type TSUBcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type TSUBcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b100011::word6) then \ \\TSUBccTV\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type TSUBccTV,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type TSUBccTV,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b100100::word6) then \ \\MULScc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type MULScc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type MULScc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b001010::word6) then \ \\UMUL\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type UMUL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type UMUL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b011010::word6) then \ \\UMULcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type UMULcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type UMULcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b001011::word6) then \ \\SMUL\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type SMUL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type SMUL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b011011::word6) then \ \\SMULcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type SMULcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type SMULcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b001110::word6) then \ \\UDIV\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type UDIV,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type UDIV,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b011110::word6) then \ \\UDIVcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type UDIVcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type UDIVcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b001111::word6) then \ \\SDIV\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type SDIV,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type SDIV,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b011111::word6) then \ \\SDIVcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type SDIVcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type SDIVcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b111100::word6) then \ \\SAVE\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (ctrl_type SAVE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (ctrl_type SAVE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b111101::word6) then \ \\RESTORE\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (ctrl_type RESTORE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (ctrl_type RESTORE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b111010::word6) then \ \\Ticc\\ let trap_cond = get_trap_cond w in let trap_imm7 = get_trap_imm7 w in if trap_cond = uint(0b1000::word4) then \ \\TA\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b0000::word4) then \ \\TN\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TN,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TN,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b1001::word4) then \ \\TNE\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TNE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TNE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b0001::word4) then \ \\TE\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b1010::word4) then \ \\TG\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TG,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TG,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b0010::word4) then \ \\TLE\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TLE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TLE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b1011::word4) then \ \\TGE\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TGE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TGE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b0011::word4) then \ \\TL\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b1100::word4) then \ \\TGU\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TGU,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TGU,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b0100::word4) then \ \\TLEU\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TLEU,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TLEU,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b1101::word4) then \ \\TCC\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TCC,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TCC,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b0101::word4) then \ \\TCS\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TCS,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TCS,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b1110::word4) then \ \\TPOS\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TPOS,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TPOS,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b0110::word4) then \ \\TNEG\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TNEG,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TNEG,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b1111::word4) then \ \\TVC\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TVC,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TVC,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b0111::word4) then \ \\TVS\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TVS,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TVS,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else Inl [invalid_trap_cond] else Inl [invalid_op3_f3_op10] else Inl [invalid_op_f3] " text \Read the word32 value from the Program Counter in the current state. Find the instruction in the memory address of the word32 value. Return a word32 value of the insturction.\ definition fetch_instruction::"('a) sparc_state \ (Exception list + word32)" where "fetch_instruction s \ \ \\pc_val\ is the 32-bit memory address of the instruction.\ let pc_val = cpu_reg_val PC s; psr_val = cpu_reg_val PSR s; s_val = get_S psr_val; asi = if s_val = 0 then word_of_int 8 else word_of_int 9 in \ \Check if \pc_val\ is aligned to 4-byte (32-bit) boundary.\ \ \That is, check if the least significant two bits of\ \ \\pc_val\ are 0s.\ if uint((AND) (0b00000000000000000000000000000011) pc_val) = 0 then \ \Get the 32-bit value from the address of \pc_val\\ \ \to the address of \pc_val+3\\ let (mem_result,n_s) = memory_read asi pc_val s in case mem_result of None \ Inl [fetch_instruction_error] |Some v \ Inr v else Inl [fetch_instruction_error] " text \Decode the word32 value of an instruction into the name of the instruction and its operands.\ definition decode_instruction::"word32 \ Exception list + instruction" where "decode_instruction w \ let this_op = get_op w in if this_op = uint(0b01::word2) then \ \Instruction format 1\ parse_instr_f1 w else if this_op = uint(0b00::word2) then \ \Instruction format 2\ parse_instr_f2 w else \ \\op = 11 0r 10\, instruction format 3\ parse_instr_f3 w " text \Get the current window from the PSR\ -definition get_curr_win::"unit \ ('a,('a::len0 window_size)) sparc_state_monad" +definition get_curr_win::"unit \ ('a,('a::len window_size)) sparc_state_monad" where "get_curr_win _ \ do curr_win \ gets (\s. (ucast (get_CWP (cpu_reg_val PSR s)))); return curr_win od" text \Operational semantics for CALL\ -definition call_instr::"instruction \ ('a::len0,unit) sparc_state_monad" +definition call_instr::"instruction \ ('a::len,unit) sparc_state_monad" where "call_instr instr \ let op_list = snd instr; mem_addr = ((ucast (get_operand_w30 (op_list!0)))::word32) << 2 in do curr_win \ get_curr_win(); pc_val \ gets (\s. (cpu_reg_val PC s)); npc_val \ gets (\s. (cpu_reg_val nPC s)); write_reg pc_val curr_win (word_of_int 15); write_cpu npc_val PC; write_cpu (pc_val + mem_addr) nPC; return () od" text\Evaluate icc based on the bits N, Z, V, C in PSR and the type of branching instruction. See Sparcv8 manual Page 178.\ definition eval_icc::"sparc_operation \ word1 \ word1 \ word1 \ word1 \ int" where "eval_icc instr_name n_val z_val v_val c_val \ if instr_name = bicc_type BNE then if z_val = 0 then 1 else 0 else if instr_name = bicc_type BE then if z_val = 1 then 1 else 0 else if instr_name = bicc_type BG then if ((OR) z_val (n_val XOR v_val)) = 0 then 1 else 0 else if instr_name = bicc_type BLE then if ((OR) z_val (n_val XOR v_val)) = 1 then 1 else 0 else if instr_name = bicc_type BGE then if (n_val XOR v_val) = 0 then 1 else 0 else if instr_name = bicc_type BL then if (n_val XOR v_val) = 1 then 1 else 0 else if instr_name = bicc_type BGU then if (c_val = 0 \ z_val = 0) then 1 else 0 else if instr_name = bicc_type BLEU then if (c_val = 1 \ z_val = 1) then 1 else 0 else if instr_name = bicc_type BCC then if c_val = 0 then 1 else 0 else if instr_name = bicc_type BCS then if c_val = 1 then 1 else 0 else if instr_name = bicc_type BNEG then if n_val = 1 then 1 else 0 else if instr_name = bicc_type BA then 1 else if instr_name = bicc_type BN then 0 else if instr_name = bicc_type BPOS then if n_val = 0 then 1 else 0 else if instr_name = bicc_type BVC then if v_val = 0 then 1 else 0 else if instr_name = bicc_type BVS then if v_val = 1 then 1 else 0 else -1 " definition branch_instr_sub1:: "sparc_operation \ ('a) sparc_state \ int" where "branch_instr_sub1 instr_name s \ let n_val = get_icc_N ((cpu_reg s) PSR); z_val = get_icc_Z ((cpu_reg s) PSR); v_val = get_icc_V ((cpu_reg s) PSR); c_val = get_icc_C ((cpu_reg s) PSR) in eval_icc instr_name n_val z_val v_val c_val" text \Operational semantics for Branching insturctions. Return exception or a bool value for annulment. If the bool value is 1, then the delay instruciton is not executed, otherwise the delay instruction is executed.\ definition branch_instr::"instruction \ ('a,unit) sparc_state_monad" where "branch_instr instr \ let instr_name = fst instr; op_list = snd instr; disp22 = get_operand_w22 (op_list!1); flaga = get_operand_flag (op_list!0) in do icc_val \ gets( \s. (branch_instr_sub1 instr_name s)); npc_val \ gets (\s. (cpu_reg_val nPC s)); pc_val \ gets (\s. (cpu_reg_val PC s)); write_cpu npc_val PC; if icc_val = 1 then do write_cpu (pc_val + (sign_ext24 (((ucast(disp22))::word24) << 2))) nPC; if (instr_name = bicc_type BA) \ (flaga = 1) then do set_annul True; return () od else return () od else \ \\icc_val = 0\\ do write_cpu (npc_val + 4) nPC; if flaga = 1 then do set_annul True; return () od else return () od od" text \Operational semantics for NOP\ definition nop_instr::"instruction \ ('a,unit) sparc_state_monad" where "nop_instr instr \ return ()" text \Operational semantics for SETHI\ -definition sethi_instr::"instruction \ ('a::len0,unit) sparc_state_monad" +definition sethi_instr::"instruction \ ('a::len,unit) sparc_state_monad" where "sethi_instr instr \ let op_list = snd instr; imm22 = get_operand_w22 (op_list!0); rd = get_operand_w5 (op_list!1) in if rd \ 0 then do curr_win \ get_curr_win(); write_reg (((ucast(imm22))::word32) << 10) curr_win rd; return () od else return () " text \ Get \operand2\ based on the flag \i\, \rs1\, \rs2\, and \simm13\. If \i = 0\ then \operand2 = r[rs2]\, else \operand2 = sign_ext13(simm13)\. \op_list\ should be \[i,rs1,rs2,\]\ or \[i,rs1,simm13,\]\. \ -definition get_operand2::"inst_operand list \ ('a::len0) sparc_state +definition get_operand2::"inst_operand list \ ('a::len) sparc_state \ virtua_address" where "get_operand2 op_list s \ let flagi = get_operand_flag (op_list!0); curr_win = ucast (get_CWP (cpu_reg_val PSR s)) in if flagi = 0 then let rs2 = get_operand_w5 (op_list!2); rs2_val = user_reg_val curr_win rs2 s in rs2_val else let ext_simm13 = sign_ext13 (get_operand_simm13 (op_list!2)) in ext_simm13 " text \ Get \operand2_val\ based on the flag \i\, \rs1\, \rs2\, and \simm13\. If \i = 0\ then \operand2_val = uint r[rs2]\, else \operand2_val = sint sign_ext13(simm13)\. \op_list\ should be \[i,rs1,rs2,\]\ or \[i,rs1,simm13,\]\. \ -definition get_operand2_val::"inst_operand list \ ('a::len0) sparc_state \ int" +definition get_operand2_val::"inst_operand list \ ('a::len) sparc_state \ int" where "get_operand2_val op_list s \ let flagi = get_operand_flag (op_list!0); curr_win = ucast (get_CWP (cpu_reg_val PSR s)) in if flagi = 0 then let rs2 = get_operand_w5 (op_list!2); rs2_val = user_reg_val curr_win rs2 s in sint rs2_val else let ext_simm13 = sign_ext13 (get_operand_simm13 (op_list!2)) in sint ext_simm13 " text \ Get the address based on the flag \i\, \rs1\, \rs2\, and \simm13\. If \i = 0\ then \addr = r[rs1] + r[rs2]\, else \addr = r[rs1] + sign_ext13(simm13)\. \op_list\ should be \[i,rs1,rs2,\]\ or \[i,rs1,simm13,\]\. \ -definition get_addr::"inst_operand list \ ('a::len0) sparc_state \ virtua_address" +definition get_addr::"inst_operand list \ ('a::len) sparc_state \ virtua_address" where "get_addr op_list s \ let rs1 = get_operand_w5 (op_list!1); curr_win = ucast (get_CWP (cpu_reg_val PSR s)); rs1_val = user_reg_val curr_win rs1 s; op2 = get_operand2 op_list s in (rs1_val + op2) " text \Operational semantics for JMPL\ -definition jmpl_instr::"instruction \ ('a::len0,unit) sparc_state_monad" +definition jmpl_instr::"instruction \ ('a::len,unit) sparc_state_monad" where "jmpl_instr instr \ let op_list = snd instr; rd = get_operand_w5 (op_list!3) in do curr_win \ get_curr_win(); jmp_addr \ gets (\s. (get_addr op_list s)); if ((AND) jmp_addr 0b00000000000000000000000000000011) \ 0 then do raise_trap mem_address_not_aligned; return () od else do rd_next_val \ gets (\s. (if rd \ 0 then (cpu_reg_val PC s) else user_reg_val curr_win rd s)); write_reg rd_next_val curr_win rd; npc_val \ gets (\s. (cpu_reg_val nPC s)); write_cpu npc_val PC; write_cpu jmp_addr nPC; return () od od" text \Operational semantics for RETT\ -definition rett_instr :: "instruction \ ('a::len0,unit) sparc_state_monad" +definition rett_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "rett_instr instr \ let op_list = snd instr in do psr_val \ gets (\s. (cpu_reg_val PSR s)); curr_win \ gets (\s. (get_CWP (cpu_reg_val PSR s))); new_cwp \ gets (\s. (word_of_int (((uint curr_win) + 1) mod NWINDOWS))); new_cwp_int \ gets (\s. ((uint curr_win) + 1) mod NWINDOWS); addr \ gets (\s. (get_addr op_list s)); et_val \ gets (\s. ((ucast (get_ET psr_val))::word1)); s_val \ gets (\s. ((ucast (get_S psr_val))::word1)); ps_val \ gets (\s. ((ucast (get_PS psr_val))::word1)); wim_val \ gets (\s. (cpu_reg_val WIM s)); npc_val \ gets (\s. (cpu_reg_val nPC s)); if et_val = 1 then if s_val = 0 then do raise_trap privileged_instruction; return () od else do raise_trap illegal_instruction; return () od else if s_val = 0 then do write_cpu_tt (0b00000011::word8); set_exe_mode False; set_err_mode True; raise_trap privileged_instruction; fail () od else if (get_WIM_bit (nat new_cwp_int) wim_val) \ 0 then do write_cpu_tt (0b00000110::word8); set_exe_mode False; set_err_mode True; raise_trap window_underflow; fail () od else if ((AND) addr (0b00000000000000000000000000000011::word32)) \ 0 then do write_cpu_tt (0b00000111::word8); set_exe_mode False; set_err_mode True; raise_trap mem_address_not_aligned; fail () od else do write_cpu npc_val PC; write_cpu addr nPC; new_psr_val \ gets (\s. (update_PSR_rett new_cwp 1 ps_val psr_val)); write_cpu new_psr_val PSR; return () od od" -definition save_retore_sub1 :: "word32 \ word5 \ word5 \ ('a::len0,unit) sparc_state_monad" +definition save_retore_sub1 :: "word32 \ word5 \ word5 \ ('a::len,unit) sparc_state_monad" where "save_retore_sub1 result new_cwp rd \ do psr_val \ gets (\s. (cpu_reg_val PSR s)); new_psr_val \ gets (\s. (update_CWP new_cwp psr_val)); write_cpu new_psr_val PSR; \ \Change \CWP\ to the new window value.\ write_reg result (ucast new_cwp) rd; \ \Write result in \rd\ of the new window.\ return () od" text \Operational semantics for SAVE and RESTORE.\ -definition save_restore_instr :: "instruction \ ('a::len0,unit) sparc_state_monad" +definition save_restore_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "save_restore_instr instr \ let instr_name = fst instr; op_list = snd instr; rd = get_operand_w5 (op_list!3) in do psr_val \ gets (\s. (cpu_reg_val PSR s)); curr_win \ get_curr_win(); wim_val \ gets (\s. (cpu_reg_val WIM s)); if instr_name = ctrl_type SAVE then do new_cwp \ gets (\s. ((word_of_int (((uint curr_win) - 1) mod NWINDOWS)))::word5); if (get_WIM_bit (unat new_cwp) wim_val) \ 0 then do raise_trap window_overflow; return () od else do result \ gets (\s. (get_addr op_list s)); \ \operands are from the old window.\ save_retore_sub1 result new_cwp rd od od else \ \\instr_name = RESTORE\\ do new_cwp \ gets (\s. ((word_of_int (((uint curr_win) + 1) mod NWINDOWS)))::word5); if (get_WIM_bit (unat new_cwp) wim_val) \ 0 then do raise_trap window_underflow; return () od else do result \ gets (\s. (get_addr op_list s)); \ \operands are from the old window.\ save_retore_sub1 result new_cwp rd od od od" definition flush_cache_line :: "word32 \ ('a,unit) sparc_state_monad" where "flush_cache_line \ undefined" definition flush_Ibuf_and_pipeline :: "word32 \ ('a,unit) sparc_state_monad" where "flush_Ibuf_and_pipeline \ undefined" text \Operational semantics for FLUSH. Flush the all the caches.\ -definition flush_instr :: "instruction \ ('a::len0,unit) sparc_state_monad" +definition flush_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "flush_instr instr \ let op_list = snd instr in do addr \ gets (\s. (get_addr op_list s)); modify (\s. (flush_cache_all s)); \<^cancel>\flush_cache_line(addr);\ \<^cancel>\flush_Ibuf_and_pipeline(addr);\ return () od" text \Operational semantics for read state register instructions. We do not consider RDASR here.\ -definition read_state_reg_instr :: "instruction \ ('a::len0,unit) sparc_state_monad" +definition read_state_reg_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "read_state_reg_instr instr \ let instr_name = fst instr; op_list = snd instr; rs1 = get_operand_w5 (op_list!0); rd = get_operand_w5 (op_list!1) in do curr_win \ get_curr_win(); psr_val \ gets (\s. (cpu_reg_val PSR s)); s_val \ gets (\s. (get_S psr_val)); if (instr_name \ {sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR} \ (instr_name = sreg_type RDASR \ privileged_ASR rs1)) \ ((ucast s_val)::word1) = 0 then do raise_trap privileged_instruction; return () od else if illegal_instruction_ASR rs1 then do raise_trap illegal_instruction; return () od else if rd \ 0 then if instr_name = sreg_type RDY then do y_val \ gets (\s. (cpu_reg_val Y s)); write_reg y_val curr_win rd; return () od else if instr_name = sreg_type RDASR then do asr_val \ gets (\s. (cpu_reg_val (ASR rs1) s)); write_reg asr_val curr_win rd; return () od else if instr_name = sreg_type RDPSR then do write_reg psr_val curr_win rd; return () od else if instr_name = sreg_type RDWIM then do wim_val \ gets (\s. (cpu_reg_val WIM s)); write_reg wim_val curr_win rd; return () od else \ \Must be \RDTBR\.\ do tbr_val \ gets (\s. (cpu_reg_val TBR s)); write_reg tbr_val curr_win rd; return () od else return () od" text \Operational semantics for write state register instructions. We do not consider WRASR here.\ -definition write_state_reg_instr :: "instruction \ ('a::len0,unit) sparc_state_monad" +definition write_state_reg_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "write_state_reg_instr instr \ let instr_name = fst instr; op_list = snd instr; rs1 = get_operand_w5 (op_list!1); rd = get_operand_w5 (op_list!3) in do curr_win \ get_curr_win(); psr_val \ gets (\s. (cpu_reg_val PSR s)); s_val \ gets (\s. (get_S psr_val)); op2 \ gets (\s. (get_operand2 op_list s)); rs1_val \ gets (\s. (user_reg_val curr_win rs1 s)); result \ gets (\s. ((XOR) rs1_val op2)); if instr_name = sreg_type WRY then do modify (\s. (delayed_pool_add (DELAYNUM, result, Y) s)); return () od else if instr_name = sreg_type WRASR then if privileged_ASR rd \ s_val = 0 then do raise_trap privileged_instruction; return () od else if illegal_instruction_ASR rd then do raise_trap illegal_instruction; return () od else do modify (\s. (delayed_pool_add (DELAYNUM, result, (ASR rd)) s)); return () od else if instr_name = sreg_type WRPSR then if s_val = 0 then do raise_trap privileged_instruction; return () od else if (uint ((ucast result)::word5)) \ NWINDOWS then do raise_trap illegal_instruction; return () od else do \ \\ET\ and \PIL\ appear to be written IMMEDIATELY w.r.t. interrupts.\ pil_val \ gets (\s. (get_PIL result)); et_val \ gets (\s. (get_ET result)); new_psr_val \ gets (\s. (update_PSR_et_pil et_val pil_val psr_val)); write_cpu new_psr_val PSR; modify (\s. (delayed_pool_add (DELAYNUM, result, PSR) s)); return () od else if instr_name = sreg_type WRWIM then if s_val = 0 then do raise_trap privileged_instruction; return () od else do \ \Don't write bits corresponding to non-existent windows.\ result_f \ gets (\s. ((result << nat (32 - NWINDOWS)) >> nat (32 - NWINDOWS))); modify (\s. (delayed_pool_add (DELAYNUM, result_f, WIM) s)); return () od else \ \Must be \WRTBR\\ if s_val = 0 then do raise_trap privileged_instruction; return () od else do \ \Only write the bits \<31:12>\ of the result to \TBR\.\ tbr_val \ gets (\s. (cpu_reg_val TBR s)); tbr_val_11_0 \ gets (\s. ((AND) tbr_val 0b00000000000000000000111111111111)); result_tmp \ gets (\s. ((AND) result 0b11111111111111111111000000000000)); result_f \ gets (\s. ((OR) tbr_val_11_0 result_tmp)); modify (\s. (delayed_pool_add (DELAYNUM, result_f, TBR) s)); return () od od" definition logical_result :: "sparc_operation \ word32 \ word32 \ word32" where "logical_result instr_name rs1_val operand2 \ if (instr_name = logic_type ANDs) \ (instr_name = logic_type ANDcc) then (AND) rs1_val operand2 else if (instr_name = logic_type ANDN) \ (instr_name = logic_type ANDNcc) then (AND) rs1_val (NOT operand2) else if (instr_name = logic_type ORs) \ (instr_name = logic_type ORcc) then (OR) rs1_val operand2 else if instr_name \ {logic_type ORN,logic_type ORNcc} then (OR) rs1_val (NOT operand2) else if instr_name \ {logic_type XORs,logic_type XORcc} then (XOR) rs1_val operand2 else \ \Must be \XNOR\ or \XNORcc\\ (XOR) rs1_val (NOT operand2) " definition logical_new_psr_val :: "word32 \ ('a) sparc_state \ word32" where "logical_new_psr_val result s \ let psr_val = cpu_reg_val PSR s; n_val = (ucast (result >> 31))::word1; z_val = if (result = 0) then 1 else 0; v_val = 0; c_val = 0 in update_PSR_icc n_val z_val v_val c_val psr_val " definition logical_instr_sub1 :: "sparc_operation \ word32 \ - ('a::len0,unit) sparc_state_monad" + ('a::len,unit) sparc_state_monad" where "logical_instr_sub1 instr_name result \ if instr_name \ {logic_type ANDcc,logic_type ANDNcc,logic_type ORcc, logic_type ORNcc,logic_type XORcc,logic_type XNORcc} then do new_psr_val \ gets (\s. (logical_new_psr_val result s)); write_cpu new_psr_val PSR; return () od else return () " text \Operational semantics for logical instructions.\ -definition logical_instr :: "instruction \ ('a::len0,unit) sparc_state_monad" +definition logical_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "logical_instr instr \ let instr_name = fst instr; op_list = snd instr; rs1 = get_operand_w5 (op_list!1); rd = get_operand_w5 (op_list!3) in do operand2 \ gets (\s. (get_operand2 op_list s)); curr_win \ get_curr_win(); rs1_val \ gets (\s. (user_reg_val curr_win rs1 s)); rd_val \ gets (\s. (user_reg_val curr_win rd s)); result \ gets (\s. (logical_result instr_name rs1_val operand2)); new_rd_val \ gets (\s. (if rd \ 0 then result else rd_val)); write_reg new_rd_val curr_win rd; logical_instr_sub1 instr_name result od" text \Operational semantics for shift instructions.\ -definition shift_instr :: "instruction \ ('a::len0,unit) sparc_state_monad" +definition shift_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "shift_instr instr \ let instr_name = fst instr; op_list = snd instr; flagi = get_operand_flag (op_list!0); rs1 = get_operand_w5 (op_list!1); rs2_shcnt = get_operand_w5 (op_list!2); rd = get_operand_w5 (op_list!3) in do curr_win \ get_curr_win(); shift_count \ gets (\s. (if flagi = 0 then ucast (user_reg_val curr_win rs2_shcnt s) else rs2_shcnt)); rs1_val \ gets (\s. (user_reg_val curr_win rs1 s)); if (instr_name = shift_type SLL) \ (rd \ 0) then do rd_val \ gets (\s. (rs1_val << (unat shift_count))); write_reg rd_val curr_win rd; return () od else if (instr_name = shift_type SRL) \ (rd \ 0) then do rd_val \ gets (\s. (rs1_val >> (unat shift_count))); write_reg rd_val curr_win rd; return () od else if (instr_name = shift_type SRA) \ (rd \ 0) then do rd_val \ gets (\s. (rs1_val >>> (unat shift_count))); write_reg rd_val curr_win rd; return () od else return () od" definition add_instr_sub1 :: "sparc_operation \ word32 \ word32 \ word32 - \ ('a::len0,unit) sparc_state_monad" + \ ('a::len,unit) sparc_state_monad" where "add_instr_sub1 instr_name result rs1_val operand2 \ if instr_name \ {arith_type ADDcc,arith_type ADDXcc} then do psr_val \ gets (\s. (cpu_reg_val PSR s)); result_31 \ gets (\s. ((ucast (result >> 31))::word1)); rs1_val_31 \ gets (\s. ((ucast (rs1_val >> 31))::word1)); operand2_31 \ gets (\s. ((ucast (operand2 >> 31))::word1)); new_n_val \ gets (\s. (result_31)); new_z_val \ gets (\s. (if result = 0 then 1::word1 else 0::word1)); new_v_val \ gets (\s. ((OR) ((AND) rs1_val_31 ((AND) operand2_31 (NOT result_31))) ((AND) (NOT rs1_val_31) ((AND) (NOT operand2_31) result_31)))); new_c_val \ gets (\s. ((OR) ((AND) rs1_val_31 operand2_31) ((AND) (NOT result_31) ((OR) rs1_val_31 operand2_31)))); new_psr_val \ gets (\s. (update_PSR_icc new_n_val new_z_val new_v_val new_c_val psr_val)); write_cpu new_psr_val PSR; return () od else return () " text \Operational semantics for add instructions. These include ADD, ADDcc, ADDX.\ -definition add_instr :: "instruction \ ('a::len0,unit) sparc_state_monad" +definition add_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "add_instr instr \ let instr_name = fst instr; op_list = snd instr; rs1 = get_operand_w5 (op_list!1); rd = get_operand_w5 (op_list!3) in do operand2 \ gets (\s. (get_operand2 op_list s)); curr_win \ get_curr_win(); rs1_val \ gets (\s. (user_reg_val curr_win rs1 s)); psr_val \ gets (\s. (cpu_reg_val PSR s)); c_val \ gets (\s. (get_icc_C psr_val)); result \ gets (\s. (if (instr_name = arith_type ADD) \ (instr_name = arith_type ADDcc) then rs1_val + operand2 else \ \Must be \ADDX\ or \ADDXcc\\ rs1_val + operand2 + (ucast c_val))); rd_val \ gets (\s. (user_reg_val curr_win rd s)); new_rd_val \ gets (\s. (if rd \ 0 then result else rd_val)); write_reg new_rd_val curr_win rd; add_instr_sub1 instr_name result rs1_val operand2 od" definition sub_instr_sub1 :: "sparc_operation \ word32 \ word32 \ word32 - \ ('a::len0,unit) sparc_state_monad" + \ ('a::len,unit) sparc_state_monad" where "sub_instr_sub1 instr_name result rs1_val operand2 \ if instr_name \ {arith_type SUBcc,arith_type SUBXcc} then do psr_val \ gets (\s. (cpu_reg_val PSR s)); result_31 \ gets (\s. ((ucast (result >> 31))::word1)); rs1_val_31 \ gets (\s. ((ucast (rs1_val >> 31))::word1)); operand2_31 \ gets (\s. ((ucast (operand2 >> 31))::word1)); new_n_val \ gets (\s. (result_31)); new_z_val \ gets (\s. (if result = 0 then 1::word1 else 0::word1)); new_v_val \ gets (\s. ((OR) ((AND) rs1_val_31 ((AND) (NOT operand2_31) (NOT result_31))) ((AND) (NOT rs1_val_31) ((AND) operand2_31 result_31)))); new_c_val \ gets (\s. ((OR) ((AND) (NOT rs1_val_31) operand2_31) ((AND) result_31 ((OR) (NOT rs1_val_31) operand2_31)))); new_psr_val \ gets (\s. (update_PSR_icc new_n_val new_z_val new_v_val new_c_val psr_val)); write_cpu new_psr_val PSR; return () od else return () " text \Operational semantics for subtract instructions. These include SUB, SUBcc, SUBX.\ -definition sub_instr :: "instruction \ ('a::len0,unit) sparc_state_monad" +definition sub_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "sub_instr instr \ let instr_name = fst instr; op_list = snd instr; rs1 = get_operand_w5 (op_list!1); rd = get_operand_w5 (op_list!3) in do operand2 \ gets (\s. (get_operand2 op_list s)); curr_win \ get_curr_win(); rs1_val \ gets (\s. (user_reg_val curr_win rs1 s)); psr_val \ gets (\s. (cpu_reg_val PSR s)); c_val \ gets (\s. (get_icc_C psr_val)); result \ gets (\s. (if (instr_name = arith_type SUB) \ (instr_name = arith_type SUBcc) then rs1_val - operand2 else \ \Must be \SUBX\ or \SUBXcc\\ rs1_val - operand2 - (ucast c_val))); rd_val \ gets (\s. (user_reg_val curr_win rd s)); new_rd_val \ gets (\s. (if rd \ 0 then result else rd_val)); write_reg new_rd_val curr_win rd; sub_instr_sub1 instr_name result rs1_val operand2 od" definition mul_instr_sub1 :: "sparc_operation \ word32 \ - ('a::len0,unit) sparc_state_monad" + ('a::len,unit) sparc_state_monad" where "mul_instr_sub1 instr_name result \ if instr_name \ {arith_type SMULcc,arith_type UMULcc} then do psr_val \ gets (\s. (cpu_reg_val PSR s)); new_n_val \ gets (\s. ((ucast (result >> 31))::word1)); new_z_val \ gets (\s. (if result = 0 then 1 else 0)); new_v_val \ gets (\s. 0); new_c_val \ gets (\s. 0); new_psr_val \ gets (\s. (update_PSR_icc new_n_val new_z_val new_v_val new_c_val psr_val)); write_cpu new_psr_val PSR; return () od else return () " text \Operational semantics for multiply instructions.\ -definition mul_instr :: "instruction \ ('a::len0,unit) sparc_state_monad" +definition mul_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "mul_instr instr \ let instr_name = fst instr; op_list = snd instr; rs1 = get_operand_w5 (op_list!1); rd = get_operand_w5 (op_list!3) in do operand2 \ gets (\s. (get_operand2 op_list s)); curr_win \ get_curr_win(); rs1_val \ gets (\s. (user_reg_val curr_win rs1 s)); psr_val \ gets (\s. (cpu_reg_val PSR s)); result0 \ gets (\s. (if instr_name \ {arith_type UMUL,arith_type UMULcc} then (word_of_int ((uint rs1_val) * (uint operand2)))::word64 else \ \Must be \SMUL\ or \SMULcc\\ (word_of_int ((sint rs1_val) * (sint operand2)))::word64)); \ \whether to use \ucast\ or \scast\ does not matter below.\ y_val \ gets (\s. ((ucast (result0 >> 32))::word32)); write_cpu y_val Y; result \ gets (\s. ((ucast result0)::word32)); rd_val \ gets (\s. (user_reg_val curr_win rd s)); new_rd_val \ gets (\s. (if rd \ 0 then result else rd_val)); write_reg new_rd_val curr_win rd; mul_instr_sub1 instr_name result od" definition div_comp_temp_64bit :: "instruction \ word64 \ virtua_address \ word64" where "div_comp_temp_64bit i y_rs1 operand2 \ if ((fst i) = arith_type UDIV) \ ((fst i) = arith_type UDIVcc) then (word_of_int ((uint y_rs1) div (uint operand2)))::word64 else \ \Must be \SDIV\ or \SDIVcc\.\ \ \Due to Isabelle's rounding method is not nearest to zero,\ \ \we have to implement division in a different way.\ let sop1 = sint y_rs1; sop2 = sint operand2; pop1 = abs sop1; pop2 = abs sop2 in if sop1 > 0 \ sop2 > 0 then (word_of_int (sop1 div sop2)) else if sop1 > 0 \ sop2 < 0 then (word_of_int (- (sop1 div pop2))) else if sop1 < 0 \ sop2 > 0 then (word_of_int (- (pop1 div sop2))) else \ \\sop1 < 0 \ sop2 < 0\\ (word_of_int (pop1 div pop2))" definition div_comp_temp_V :: "instruction \ word32 \ word33 \ word1" where "div_comp_temp_V i w32 w33 \ if ((fst i) = arith_type UDIV) \ ((fst i) = arith_type UDIVcc) then if w32 = 0 then 0 else 1 else \ \Must be \SDIV\ or \SDIVcc\.\ if (w33 = 0) \ (w33 = (0b111111111111111111111111111111111::word33)) then 0 else 1" definition div_comp_result :: "instruction \ word1 \ word64 \ word32" where "div_comp_result i temp_V temp_64bit \ if temp_V = 1 then if ((fst i) = arith_type UDIV) \ ((fst i) = arith_type UDIVcc) then (0b11111111111111111111111111111111::word32) else if (fst i) \ {arith_type SDIV,arith_type SDIVcc} then if temp_64bit > 0 then (0b01111111111111111111111111111111::word32) else ((word_of_int (0 - (uint (0b10000000000000000000000000000000::word32))))::word32) else ((ucast temp_64bit)::word32) else ((ucast temp_64bit)::word32)" definition div_write_new_val :: "instruction \ word32 \ word1 \ - ('a::len0,unit) sparc_state_monad" + ('a::len,unit) sparc_state_monad" where "div_write_new_val i result temp_V \ if (fst i) \ {arith_type UDIVcc,arith_type SDIVcc} then do psr_val \ gets (\s. (cpu_reg_val PSR s)); new_n_val \ gets (\s. ((ucast (result >> 31))::word1)); new_z_val \ gets (\s. (if result = 0 then 1 else 0)); new_v_val \ gets (\s. temp_V); new_c_val \ gets (\s. 0); new_psr_val \ gets (\s. (update_PSR_icc new_n_val new_z_val new_v_val new_c_val psr_val)); write_cpu new_psr_val PSR; return () od else return ()" definition div_comp :: "instruction \ word5 \ word5 \ virtua_address \ - ('a::len0,unit) sparc_state_monad" + ('a::len,unit) sparc_state_monad" where "div_comp instr rs1 rd operand2 \ do curr_win \ get_curr_win(); rs1_val \ gets (\s. (user_reg_val curr_win rs1 s)); y_val \ gets (\s. (cpu_reg_val Y s)); y_rs1 \ gets (\s. ((word_cat y_val rs1_val)::word64)); temp_64bit \ gets (\s. (div_comp_temp_64bit instr y_rs1 operand2)); \<^cancel>\result \ gets (\s. ((ucast temp_64bit)::word32));\ temp_high32 \ gets (\s. ((ucast (temp_64bit >> 32))::word32)); temp_high33 \ gets (\s. ((ucast (temp_64bit >> 31))::word33)); temp_V \ gets (\s. (div_comp_temp_V instr temp_high32 temp_high33)); result \ gets (\s. (div_comp_result instr temp_V temp_64bit)); rd_val \ gets (\s. (user_reg_val curr_win rd s)); new_rd_val \ gets (\s. (if rd \ 0 then result else rd_val)); write_reg new_rd_val curr_win rd; div_write_new_val instr result temp_V od" text \Operational semantics for divide instructions.\ -definition div_instr :: "instruction \ ('a::len0,unit) sparc_state_monad" +definition div_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "div_instr instr \ let instr_name = fst instr; op_list = snd instr; rs1 = get_operand_w5 (op_list!1); rd = get_operand_w5 (op_list!3) in do operand2 \ gets (\s. (get_operand2 op_list s)); if (uint operand2) = 0 then do raise_trap division_by_zero; return () od else div_comp instr rs1 rd operand2 od" definition ld_word0 :: "instruction \ word32 \ virtua_address \ word32" where "ld_word0 instr data_word address \ if (fst instr) \ {load_store_type LDSB,load_store_type LDUB, load_store_type LDUBA,load_store_type LDSBA} then let byte = if (uint ((ucast address)::word2)) = 0 then (ucast (data_word >> 24))::word8 else if (uint ((ucast address)::word2)) = 1 then (ucast (data_word >> 16))::word8 else if (uint ((ucast address)::word2)) = 2 then (ucast (data_word >> 8))::word8 else \ \Must be 3.\ (ucast data_word)::word8 in if (fst instr) = load_store_type LDSB \ (fst instr) = load_store_type LDSBA then sign_ext8 byte else zero_ext8 byte else if (fst instr) = load_store_type LDUH \ (fst instr) = load_store_type LDSH \ (fst instr) = load_store_type LDSHA \ (fst instr) = load_store_type LDUHA then let halfword = if (uint ((ucast address)::word2)) = 0 then (ucast (data_word >> 16))::word16 else \ \Must be 2.\ (ucast data_word)::word16 in if (fst instr) = load_store_type LDSH \ (fst instr) = load_store_type LDSHA then sign_ext16 halfword else zero_ext16 halfword else \ \Must be LDD\ data_word " definition ld_asi :: "instruction \ word1 \ asi_type" where "ld_asi instr s_val \ if (fst instr) \ {load_store_type LDD,load_store_type LD,load_store_type LDUH, load_store_type LDSB,load_store_type LDUB,load_store_type LDSH} then if s_val = 0 then (word_of_int 10)::asi_type else (word_of_int 11)::asi_type else \ \Must be \LDA\, \LDUBA\, \LDSBA\, \LDSHA\, \LDUHA\, or \LDDA\.\ get_operand_asi ((snd instr)!3) " definition load_sub2 :: "virtua_address \ asi_type \ word5 \ - ('a::len0) window_size \ word32 \ ('a,unit) sparc_state_monad" + ('a::len) window_size \ word32 \ ('a,unit) sparc_state_monad" where "load_sub2 address asi rd curr_win word0 \ do write_reg word0 curr_win ((AND) rd 0b11110); (result1,new_state1) \ gets (\s. (memory_read asi (address + 4) s)); if result1 = None then do raise_trap data_access_exception; return () od else do word1 \ gets (\s. (case result1 of Some v \ v)); modify (\s. (new_state1)); write_reg word1 curr_win ((OR) rd 1); return () od od" -definition load_sub3 :: "instruction \ ('a::len0) window_size \ +definition load_sub3 :: "instruction \ ('a::len) window_size \ word5 \ asi_type \ virtua_address \ - ('a::len0,unit) sparc_state_monad" + ('a::len,unit) sparc_state_monad" where "load_sub3 instr curr_win rd asi address \ do (result,new_state) \ gets (\s. (memory_read asi address s)); if result = None then do raise_trap data_access_exception; return () od else do data_word \ gets (\s. (case result of Some v \ v)); modify (\s. (new_state)); word0 \ gets (\s. (ld_word0 instr data_word address)); if rd \ 0 \ (fst instr) \ {load_store_type LD,load_store_type LDA, load_store_type LDUH,load_store_type LDSB,load_store_type LDUB, load_store_type LDUBA,load_store_type LDSH,load_store_type LDSHA, load_store_type LDUHA,load_store_type LDSBA} then do write_reg word0 curr_win rd; return () od else \ \Must be \LDD\ or \LDDA\\ load_sub2 address asi rd curr_win word0 od od" definition load_sub1 :: "instruction \ word5 \ word1 \ - ('a::len0,unit) sparc_state_monad" + ('a::len,unit) sparc_state_monad" where "load_sub1 instr rd s_val \ do curr_win \ get_curr_win(); address \ gets (\s. (get_addr (snd instr) s)); asi \ gets (\s. (ld_asi instr s_val)); if (((fst instr) = load_store_type LDD \ (fst instr) = load_store_type LDDA) \ ((ucast address)::word3) \ 0) \ ((fst instr) \ {load_store_type LD,load_store_type LDA} \ ((ucast address)::word2) \ 0) \ (((fst instr) = load_store_type LDUH \ (fst instr) = load_store_type LDUHA \ (fst instr) = load_store_type LDSH \ (fst instr) = load_store_type LDSHA) \ ((ucast address)::word1) \ 0) then do raise_trap mem_address_not_aligned; return () od else load_sub3 instr curr_win rd asi address od" text \Operational semantics for Load instructions.\ -definition load_instr :: "instruction \ ('a::len0,unit) sparc_state_monad" +definition load_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "load_instr instr \ let instr_name = fst instr; op_list = snd instr; flagi = get_operand_flag (op_list!0); rd = if instr_name \ {load_store_type LDUBA,load_store_type LDA, load_store_type LDSBA,load_store_type LDSHA, load_store_type LDSHA,load_store_type LDDA} then \ \\rd\ is member 4\ get_operand_w5 (op_list!4) else \ \\rd\ is member 3\ get_operand_w5 (op_list!3) in do psr_val \ gets (\s. (cpu_reg_val PSR s)); s_val \ gets (\s. (get_S psr_val)); if instr_name \ {load_store_type LDA,load_store_type LDUBA, load_store_type LDSBA,load_store_type LDSHA, load_store_type LDUHA,load_store_type LDDA} \ s_val = 0 then do raise_trap privileged_instruction; return () od else if instr_name \ {load_store_type LDA,load_store_type LDUBA, load_store_type LDSBA,load_store_type LDSHA,load_store_type LDUHA, load_store_type LDDA} \ flagi = 1 then do raise_trap illegal_instruction; return () od else load_sub1 instr rd s_val od" definition st_asi :: "instruction \ word1 \ asi_type" where "st_asi instr s_val \ if (fst instr) \ {load_store_type STD,load_store_type ST, load_store_type STH,load_store_type STB} then if s_val = 0 then (word_of_int 10)::asi_type else (word_of_int 11)::asi_type else \ \Must be \STA\, \STBA\, \STHA\, \STDA\.\ get_operand_asi ((snd instr)!3) " definition st_byte_mask :: "instruction \ virtua_address \ word4" where "st_byte_mask instr address \ if (fst instr) \ {load_store_type STD,load_store_type ST, load_store_type STA,load_store_type STDA} then (0b1111::word4) else if (fst instr) \ {load_store_type STH,load_store_type STHA} then if ((ucast address)::word2) = 0 then (0b1100::word4) else \ \Must be 2.\ (0b0011::word4) else \ \Must be \STB\ or \STBA\.\ if ((ucast address)::word2) = 0 then (0b1000::word4) else if ((ucast address)::word2) = 1 then (0b0100::word4) else if ((ucast address)::word2) = 2 then (0b0010::word4) else \ \Must be 3.\ (0b0001::word4) " -definition st_data0 :: "instruction \ ('a::len0) window_size \ +definition st_data0 :: "instruction \ ('a::len) window_size \ word5 \ virtua_address \ ('a) sparc_state \ reg_type" where "st_data0 instr curr_win rd address s \ if (fst instr) \ {load_store_type STD,load_store_type STDA} then user_reg_val curr_win ((AND) rd 0b11110) s else if (fst instr) \ {load_store_type ST,load_store_type STA} then user_reg_val curr_win rd s else if (fst instr) \ {load_store_type STH,load_store_type STHA} then if ((ucast address)::word2) = 0 then (user_reg_val curr_win rd s) << 16 else \ \Must be 2.\ user_reg_val curr_win rd s else \ \Must be \STB\ or \STBA\.\ if ((ucast address)::word2) = 0 then (user_reg_val curr_win rd s) << 24 else if ((ucast address)::word2) = 1 then (user_reg_val curr_win rd s) << 16 else if ((ucast address)::word2) = 2 then (user_reg_val curr_win rd s) << 8 else \ \Must be 3.\ user_reg_val curr_win rd s " -definition store_sub2 :: "instruction \ ('a::len0) window_size \ +definition store_sub2 :: "instruction \ ('a::len) window_size \ word5 \ asi_type \ virtua_address \ - ('a::len0,unit) sparc_state_monad" + ('a::len,unit) sparc_state_monad" where "store_sub2 instr curr_win rd asi address \ do byte_mask \ gets (\s. (st_byte_mask instr address)); data0 \ gets (\s. (st_data0 instr curr_win rd address s)); result0 \ gets (\s. (memory_write asi address byte_mask data0 s)); if result0 = None then do raise_trap data_access_exception; return () od else do new_state \ gets (\s. (case result0 of Some v \ v)); modify (\s. (new_state)); if (fst instr) \ {load_store_type STD,load_store_type STDA} then do data1 \ gets (\s. (user_reg_val curr_win ((OR) rd 0b00001) s)); result1 \ gets (\s. (memory_write asi (address + 4) (0b1111::word4) data1 s)); if result1 = None then do raise_trap data_access_exception; return () od else do new_state1 \ gets (\s. (case result1 of Some v \ v)); modify (\s. (new_state1)); return () od od else return () od od" definition store_sub1 :: "instruction \ word5 \ word1 \ - ('a::len0,unit) sparc_state_monad" + ('a::len,unit) sparc_state_monad" where "store_sub1 instr rd s_val \ do curr_win \ get_curr_win(); address \ gets (\s. (get_addr (snd instr) s)); asi \ gets (\s. (st_asi instr s_val)); \ \The following code is intentionally long to match the definitions in SPARCv8.\ if ((fst instr) = load_store_type STH \ (fst instr) = load_store_type STHA) \ ((ucast address)::word1) \ 0 then do raise_trap mem_address_not_aligned; return () od else if (fst instr) \ {load_store_type ST,load_store_type STA} \ ((ucast address)::word2) \ 0 then do raise_trap mem_address_not_aligned; return () od else if (fst instr) \ {load_store_type STD,load_store_type STDA} \ ((ucast address)::word3) \ 0 then do raise_trap mem_address_not_aligned; return () od else store_sub2 instr curr_win rd asi address od" text \Operational semantics for Store instructions.\ definition store_instr :: "instruction \ - ('a::len0,unit) sparc_state_monad" + ('a::len,unit) sparc_state_monad" where "store_instr instr \ let instr_name = fst instr; op_list = snd instr; flagi = get_operand_flag (op_list!0); rd = if instr_name \ {load_store_type STA,load_store_type STBA, load_store_type STHA,load_store_type STDA} then \ \\rd\ is member 4\ get_operand_w5 (op_list!4) else \ \\rd\ is member 3\ get_operand_w5 (op_list!3) in do psr_val \ gets (\s. (cpu_reg_val PSR s)); s_val \ gets (\s. (get_S psr_val)); if instr_name \ {load_store_type STA,load_store_type STDA, load_store_type STHA,load_store_type STBA} \ s_val = 0 then do raise_trap privileged_instruction; return () od else if instr_name \ {load_store_type STA,load_store_type STDA, load_store_type STHA,load_store_type STBA} \ flagi = 1 then do raise_trap illegal_instruction; return () od else store_sub1 instr rd s_val od" text \The instructions below are not used by Xtratum and they are not tested.\ definition ldst_asi :: "instruction \ word1 \ asi_type" where "ldst_asi instr s_val \ if (fst instr) \ {load_store_type LDSTUB} then if s_val = 0 then (word_of_int 10)::asi_type else (word_of_int 11)::asi_type else \ \Must be \LDSTUBA\.\ get_operand_asi ((snd instr)!3) " definition ldst_word0 :: "instruction \ word32 \ virtua_address \ word32" where "ldst_word0 instr data_word address \ let byte = if (uint ((ucast address)::word2)) = 0 then (ucast (data_word >> 24))::word8 else if (uint ((ucast address)::word2)) = 1 then (ucast (data_word >> 16))::word8 else if (uint ((ucast address)::word2)) = 2 then (ucast (data_word >> 8))::word8 else \ \Must be 3.\ (ucast data_word)::word8 in zero_ext8 byte " definition ldst_byte_mask :: "instruction \ virtua_address \ word4" where "ldst_byte_mask instr address \ if ((ucast address)::word2) = 0 then (0b1000::word4) else if ((ucast address)::word2) = 1 then (0b0100::word4) else if ((ucast address)::word2) = 2 then (0b0010::word4) else \ \Must be 3.\ (0b0001::word4) " definition load_store_sub1 :: "instruction \ word5 \ word1 \ - ('a::len0,unit) sparc_state_monad" + ('a::len,unit) sparc_state_monad" where "load_store_sub1 instr rd s_val \ do curr_win \ get_curr_win(); address \ gets (\s. (get_addr (snd instr) s)); asi \ gets (\s. (ldst_asi instr s_val)); \ \wait for locks to be lifted.\ \ \an implementation actually need only block when another \LDSTUB\ or \SWAP\\ \ \is pending on the same byte in memory as the one addressed by this \LDSTUB\\ \ \Should wait when \block_type = 1 \ block_word = 1\\ \ \until another processes write both to be 0.\ \ \We implement this as setting \pc\ as \npc\ when the instruction\ \ \is blocked. This way, in the next iteration, we will still execution\ \ \the current instruction.\ block_byte \ gets (\s. (pb_block_ldst_byte_val address s)); block_word \ gets (\s. (pb_block_ldst_word_val address s)); if block_byte \ block_word then do pc_val \ gets (\s. (cpu_reg_val PC s)); write_cpu pc_val nPC; return () od else do modify (\s. (pb_block_ldst_byte_mod address True s)); (result,new_state) \ gets (\s. (memory_read asi address s)); if result = None then do raise_trap data_access_exception; return () od else do data_word \ gets (\s. (case result of Some v \ v)); modify (\s. (new_state)); byte_mask \ gets (\s. (ldst_byte_mask instr address)); data0 \ gets (\s. (0b11111111111111111111111111111111::word32)); result0 \ gets (\s. (memory_write asi address byte_mask data0 s)); modify (\s. (pb_block_ldst_byte_mod address False s)); if result0 = None then do raise_trap data_access_exception; return () od else do new_state1 \ gets (\s. (case result0 of Some v \ v)); modify (\s. (new_state1)); word0 \ gets (\s. (ldst_word0 instr data_word address)); if rd \ 0 then do write_reg word0 curr_win rd; return () od else return () od od od od" text \Operational semantics for atomic load-store.\ -definition load_store_instr :: "instruction \ ('a::len0,unit) sparc_state_monad" +definition load_store_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "load_store_instr instr \ let instr_name = fst instr; op_list = snd instr; flagi = get_operand_flag (op_list!0); rd = if instr_name \ {load_store_type LDSTUBA} then \ \\rd\ is member 4\ get_operand_w5 (op_list!4) else \ \\rd\ is member 3\ get_operand_w5 (op_list!3) in do psr_val \ gets (\s. (cpu_reg_val PSR s)); s_val \ gets (\s. (get_S psr_val)); if instr_name \ {load_store_type LDSTUBA} \ s_val = 0 then do raise_trap privileged_instruction; return () od else if instr_name \ {load_store_type LDSTUBA} \ flagi = 1 then do raise_trap illegal_instruction; return () od else load_store_sub1 instr rd s_val od" definition swap_sub1 :: "instruction \ word5 \ word1 \ - ('a::len0,unit) sparc_state_monad" + ('a::len,unit) sparc_state_monad" where "swap_sub1 instr rd s_val \ do curr_win \ get_curr_win(); address \ gets (\s. (get_addr (snd instr) s)); asi \ gets (\s. (ldst_asi instr s_val)); temp \ gets (\s. (user_reg_val curr_win rd s)); \ \wait for locks to be lifted.\ \ \an implementation actually need only block when another \LDSTUB\ or \SWAP\\ \ \is pending on the same byte in memory as the one addressed by this \LDSTUB\\ \ \Should wait when \block_type = 1 \ block_word = 1\\ \ \until another processes write both to be 0.\ \ \We implement this as setting \pc\ as \npc\ when the instruction\ \ \is blocked. This way, in the next iteration, we will still execution\ \ \the current instruction.\ block_byte \ gets (\s. (pb_block_ldst_byte_val address s)); block_word \ gets (\s. (pb_block_ldst_word_val address s)); if block_byte \ block_word then do pc_val \ gets (\s. (cpu_reg_val PC s)); write_cpu pc_val nPC; return () od else do modify (\s. (pb_block_ldst_word_mod address True s)); (result,new_state) \ gets (\s. (memory_read asi address s)); if result = None then do raise_trap data_access_exception; return () od else do word \ gets (\s. (case result of Some v \ v)); modify (\s. (new_state)); byte_mask \ gets (\s. (0b1111::word4)); result0 \ gets (\s. (memory_write asi address byte_mask temp s)); modify (\s. (pb_block_ldst_word_mod address False s)); if result0 = None then do raise_trap data_access_exception; return () od else do new_state1 \ gets (\s. (case result0 of Some v \ v)); modify (\s. (new_state1)); if rd \ 0 then do write_reg word curr_win rd; return () od else return () od od od od" text \Operational semantics for swap.\ -definition swap_instr :: "instruction \ ('a::len0,unit) sparc_state_monad" +definition swap_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "swap_instr instr \ let instr_name = fst instr; op_list = snd instr; flagi = get_operand_flag (op_list!0); rd = if instr_name \ {load_store_type SWAPA} then \ \\rd\ is member 4\ get_operand_w5 (op_list!4) else \ \\rd\ is member 3\ get_operand_w5 (op_list!3) in do psr_val \ gets (\s. (cpu_reg_val PSR s)); s_val \ gets (\s. (get_S psr_val)); if instr_name \ {load_store_type SWAPA} \ s_val = 0 then do raise_trap privileged_instruction; return () od else if instr_name \ {load_store_type SWAPA} \ flagi = 1 then do raise_trap illegal_instruction; return () od else swap_sub1 instr rd s_val od" definition bit2_zero :: "word2 \ word1" where "bit2_zero w2 \ if w2 \ 0 then 1 else 0" text \Operational semantics for tagged add instructions.\ -definition tadd_instr :: "instruction \ ('a::len0,unit) sparc_state_monad" +definition tadd_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "tadd_instr instr \ let instr_name = fst instr; op_list = snd instr; rs1 = get_operand_w5 (op_list!1); rd = get_operand_w5 (op_list!3) in do operand2 \ gets (\s. (get_operand2 op_list s)); curr_win \ get_curr_win(); rs1_val \ gets (\s. (user_reg_val curr_win rs1 s)); psr_val \ gets (\s. (cpu_reg_val PSR s)); c_val \ gets (\s. (get_icc_C psr_val)); result \ gets (\s. (rs1_val + operand2)); result_31 \ gets (\s. ((ucast (result >> 31))::word1)); rs1_val_31 \ gets (\s. ((ucast (rs1_val >> 31))::word1)); operand2_31 \ gets (\s. ((ucast (operand2 >> 31))::word1)); rs1_val_2 \ gets (\s. (bit2_zero ((ucast rs1_val)::word2))); operand2_2 \ gets (\s. (bit2_zero ((ucast operand2)::word2))); temp_V \ gets (\s. ((OR) ((OR) ((AND) rs1_val_31 ((AND) operand2_31 (NOT result_31))) ((AND) (NOT rs1_val_31) ((AND) (NOT operand2_31) result_31))) ((OR) rs1_val_2 operand2_2))); if instr_name = arith_type TADDccTV \ temp_V = 1 then do raise_trap tag_overflow; return () od else do rd_val \ gets (\s. (user_reg_val curr_win rd s)); new_rd_val \ gets (\s. (if rd \ 0 then result else rd_val)); write_reg new_rd_val curr_win rd; new_n_val \ gets (\s. (result_31)); new_z_val \ gets (\s. (if result = 0 then 1::word1 else 0::word1)); new_v_val \ gets (\s. temp_V); new_c_val \ gets (\s. ((OR) ((AND) rs1_val_31 operand2_31) ((AND) (NOT result_31) ((OR) rs1_val_31 operand2_31)))); new_psr_val \ gets (\s. (update_PSR_icc new_n_val new_z_val new_v_val new_c_val psr_val)); write_cpu new_psr_val PSR; rd_val \ gets (\s. (user_reg_val curr_win rd s)); new_rd_val \ gets (\s. (if rd \ 0 then result else rd_val)); write_reg new_rd_val curr_win rd; return () od od" text \Operational semantics for tagged add instructions.\ -definition tsub_instr :: "instruction \ ('a::len0,unit) sparc_state_monad" +definition tsub_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "tsub_instr instr \ let instr_name = fst instr; op_list = snd instr; rs1 = get_operand_w5 (op_list!1); rd = get_operand_w5 (op_list!3) in do operand2 \ gets (\s. (get_operand2 op_list s)); curr_win \ get_curr_win(); rs1_val \ gets (\s. (user_reg_val curr_win rs1 s)); psr_val \ gets (\s. (cpu_reg_val PSR s)); c_val \ gets (\s. (get_icc_C psr_val)); result \ gets (\s. (rs1_val - operand2)); result_31 \ gets (\s. ((ucast (result >> 31))::word1)); rs1_val_31 \ gets (\s. ((ucast (rs1_val >> 31))::word1)); operand2_31 \ gets (\s. ((ucast (operand2 >> 31))::word1)); rs1_val_2 \ gets (\s. (bit2_zero ((ucast rs1_val)::word2))); operand2_2 \ gets (\s. (bit2_zero ((ucast operand2)::word2))); temp_V \ gets (\s. ((OR) ((OR) ((AND) rs1_val_31 ((AND) operand2_31 (NOT result_31))) ((AND) (NOT rs1_val_31) ((AND) (NOT operand2_31) result_31))) ((OR) rs1_val_2 operand2_2))); if instr_name = arith_type TSUBccTV \ temp_V = 1 then do raise_trap tag_overflow; return () od else do rd_val \ gets (\s. (user_reg_val curr_win rd s)); new_rd_val \ gets (\s. (if rd \ 0 then result else rd_val)); write_reg new_rd_val curr_win rd; new_n_val \ gets (\s. (result_31)); new_z_val \ gets (\s. (if result = 0 then 1::word1 else 0::word1)); new_v_val \ gets (\s. temp_V); new_c_val \ gets (\s. ((OR) ((AND) rs1_val_31 operand2_31) ((AND) (NOT result_31) ((OR) rs1_val_31 operand2_31)))); new_psr_val \ gets (\s. (update_PSR_icc new_n_val new_z_val new_v_val new_c_val psr_val)); write_cpu new_psr_val PSR; rd_val \ gets (\s. (user_reg_val curr_win rd s)); new_rd_val \ gets (\s. (if rd \ 0 then result else rd_val)); write_reg new_rd_val curr_win rd; return () od od" -definition muls_op2 :: "inst_operand list \ ('a::len0) sparc_state \ word32" +definition muls_op2 :: "inst_operand list \ ('a::len) sparc_state \ word32" where "muls_op2 op_list s \ let y_val = cpu_reg_val Y s in if ((ucast y_val)::word1) = 0 then 0 else get_operand2 op_list s " text \Operational semantics for multiply step instruction.\ -definition muls_instr :: "instruction \ ('a::len0,unit) sparc_state_monad" +definition muls_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "muls_instr instr \ let instr_name = fst instr; op_list = snd instr; rs1 = get_operand_w5 (op_list!1); rd = get_operand_w5 (op_list!3) in do curr_win \ get_curr_win(); rs1_val \ gets (\s. (user_reg_val curr_win rs1 s)); psr_val \ gets (\s. (cpu_reg_val PSR s)); n_val \ gets (\s. (get_icc_N psr_val)); v_val \ gets (\s. (get_icc_V psr_val)); c_val \ gets (\s. (get_icc_C psr_val)); y_val \ gets (\s. (cpu_reg_val Y s)); operand1 \ gets (\s. (word_cat ((XOR) n_val v_val) ((ucast (rs1_val >> 1))::word31))); operand2 \ gets (\s. (muls_op2 op_list s)); result \ gets (\s. (operand1 + operand2)); new_y_val \ gets (\s. (word_cat ((ucast rs1_val)::word1) ((ucast (y_val >> 1))::word31))); write_cpu new_y_val Y; rd_val \ gets (\s. (user_reg_val curr_win rd s)); new_rd_val \ gets (\s. (if rd \ 0 then result else rd_val)); write_reg new_rd_val curr_win rd; result_31 \ gets (\s. ((ucast (result >> 31))::word1)); operand1_31 \ gets (\s. ((ucast (operand1 >> 31))::word1)); operand2_31 \ gets (\s. ((ucast (operand2 >> 31))::word1)); new_n_val \ gets (\s. (result_31)); new_z_val \ gets (\s. (if result = 0 then 1::word1 else 0::word1)); new_v_val \ gets (\s. ((OR) ((AND) operand1_31 ((AND) operand2_31 (NOT result_31))) ((AND) (NOT operand1_31) ((AND) (NOT operand2_31) result_31)))); new_c_val \ gets (\s. ((OR) ((AND) operand1_31 operand2_31) ((AND) (NOT result_31) ((OR) operand1_31 operand2_31)))); new_psr_val \ gets (\s. (update_PSR_icc new_n_val new_z_val new_v_val new_c_val psr_val)); write_cpu new_psr_val PSR; return () od" text\Evaluate icc based on the bits N, Z, V, C in PSR and the type of ticc instruction. See Sparcv8 manual Page 182.\ definition trap_eval_icc::"sparc_operation \ word1 \ word1 \ word1 \ word1 \ int" where "trap_eval_icc instr_name n_val z_val v_val c_val \ if instr_name = ticc_type TNE then if z_val = 0 then 1 else 0 else if instr_name = ticc_type TE then if z_val = 1 then 1 else 0 else if instr_name = ticc_type TG then if ((OR) z_val (n_val XOR v_val)) = 0 then 1 else 0 else if instr_name = ticc_type TLE then if ((OR) z_val (n_val XOR v_val)) = 1 then 1 else 0 else if instr_name = ticc_type TGE then if (n_val XOR v_val) = 0 then 1 else 0 else if instr_name = ticc_type TL then if (n_val XOR v_val) = 1 then 1 else 0 else if instr_name = ticc_type TGU then if (c_val = 0 \ z_val = 0) then 1 else 0 else if instr_name = ticc_type TLEU then if (c_val = 1 \ z_val = 1) then 1 else 0 else if instr_name = ticc_type TCC then if c_val = 0 then 1 else 0 else if instr_name = ticc_type TCS then if c_val = 1 then 1 else 0 else if instr_name = ticc_type TPOS then if n_val = 0 then 1 else 0 else if instr_name = ticc_type TNEG then if n_val = 1 then 1 else 0 else if instr_name = ticc_type TVC then if v_val = 0 then 1 else 0 else if instr_name = ticc_type TVS then if v_val = 1 then 1 else 0 else if instr_name = ticc_type TA then 1 else if instr_name = ticc_type TN then 0 else -1 " text \ Get \operand2\ for \ticc\ based on the flag \i\, \rs1\, \rs2\, and \trap_imm7\. If \i = 0\ then \operand2 = r[rs2]\, else \operand2 = sign_ext7(trap_imm7)\. \op_list\ should be \[i,rs1,rs2]\ or \[i,rs1,trap_imm7]\. \ -definition get_trap_op2::"inst_operand list \ ('a::len0) sparc_state +definition get_trap_op2::"inst_operand list \ ('a::len) sparc_state \ virtua_address" where "get_trap_op2 op_list s \ let flagi = get_operand_flag (op_list!0); curr_win = ucast (get_CWP (cpu_reg_val PSR s)) in if flagi = 0 then let rs2 = get_operand_w5 (op_list!2); rs2_val = user_reg_val curr_win rs2 s in rs2_val else let ext_simm7 = sign_ext7 (get_operand_imm7 (op_list!2)) in ext_simm7 " text \Operational semantics for Ticc insturctions.\ definition ticc_instr::"instruction \ - ('a::len0,unit) sparc_state_monad" + ('a::len,unit) sparc_state_monad" where "ticc_instr instr \ let instr_name = fst instr; op_list = snd instr; rs1 = get_operand_w5 (op_list!1) in do n_val \ gets (\s. get_icc_N ((cpu_reg s) PSR)); z_val \ gets (\s. get_icc_Z ((cpu_reg s) PSR)); v_val \ gets (\s. get_icc_V ((cpu_reg s) PSR)); c_val \ gets (\s. get_icc_C ((cpu_reg s) PSR)); icc_val \ gets(\s. (trap_eval_icc instr_name n_val z_val v_val c_val)); curr_win \ get_curr_win(); rs1_val \ gets (\s. (user_reg_val curr_win rs1 s)); trap_number \ gets (\s. (rs1_val + (get_trap_op2 op_list s))); npc_val \ gets (\s. (cpu_reg_val nPC s)); pc_val \ gets (\s. (cpu_reg_val PC s)); if icc_val = 1 then do raise_trap trap_instruction; trap_number7 \ gets (\s. ((ucast trap_number)::word7)); modify (\s. (ticc_trap_type_mod trap_number7 s)); return () od else \ \\icc_val = 0\\ do write_cpu npc_val PC; write_cpu (npc_val + 4) nPC; return () od od" text \Operational semantics for store barrier.\ -definition store_barrier_instr::"instruction \ ('a::len0,unit) sparc_state_monad" +definition store_barrier_instr::"instruction \ ('a::len,unit) sparc_state_monad" where "store_barrier_instr instr \ do modify (\s. (store_barrier_pending_mod True s)); return () od" end diff --git a/thys/SPARCv8/SparcModel_MMU/Sparc_Properties.thy b/thys/SPARCv8/SparcModel_MMU/Sparc_Properties.thy --- a/thys/SPARCv8/SparcModel_MMU/Sparc_Properties.thy +++ b/thys/SPARCv8/SparcModel_MMU/Sparc_Properties.thy @@ -1,8565 +1,8565 @@ (* * Copyright 2016, NTU * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * Author: Zhe Hou. *) theory Sparc_Properties imports Main Sparc_Execution begin (*********************************************************************) section\Single step theorem\ (*********************************************************************) text \The following shows that, if the pre-state satisfies certain conditions called \good_context\, there must be a defined post-state after a single step execution.\ method save_restore_proof = ((simp add: save_restore_instr_def), (simp add: Let_def simpler_gets_def bind_def h1_def h2_def), (simp add: case_prod_unfold), (simp add: raise_trap_def simpler_modify_def), (simp add: simpler_gets_def bind_def h1_def h2_def), (simp add: save_retore_sub1_def), (simp add: write_cpu_def simpler_modify_def), (simp add: write_reg_def simpler_modify_def), (simp add: get_curr_win_def), (simp add: simpler_gets_def bind_def h1_def h2_def)) method select_trap_proof0 = ((simp add: select_trap_def exec_gets return_def), (simp add: DetMonad.bind_def h1_def h2_def simpler_modify_def), (simp add: write_cpu_tt_def write_cpu_def), (simp add: DetMonad.bind_def h1_def h2_def simpler_modify_def), (simp add: return_def simpler_gets_def)) method select_trap_proof1 = ((simp add: select_trap_def exec_gets return_def), (simp add: DetMonad.bind_def h1_def h2_def simpler_modify_def), (simp add: write_cpu_tt_def write_cpu_def), (simp add: DetMonad.bind_def h1_def h2_def simpler_modify_def), (simp add: return_def simpler_gets_def), (simp add: emp_trap_set_def err_mode_val_def cpu_reg_mod_def)) method dispatch_instr_proof1 = ((simp add: dispatch_instruction_def), (simp add: simpler_gets_def bind_def h1_def h2_def), (simp add: Let_def)) method exe_proof_to_decode = ((simp add: execute_instruction_def), (simp add: exec_gets bind_def h1_def h2_def Let_def return_def), clarsimp, (simp add: simpler_gets_def bind_def h1_def h2_def Let_def simpler_modify_def), (simp add: return_def)) method exe_proof_dispatch_rett = ((simp add: dispatch_instruction_def), (simp add: simpler_gets_def bind_def h1_def h2_def Let_def), (simp add: rett_instr_def), (simp add: simpler_gets_def bind_def h1_def h2_def Let_def)) lemma write_cpu_result: "snd (write_cpu w r s) = False" by (simp add: write_cpu_def simpler_modify_def) lemma set_annul_result: "snd (set_annul b s) = False" by (simp add: set_annul_def simpler_modify_def) lemma raise_trap_result : "snd (raise_trap t s) = False" by (simp add: raise_trap_def simpler_modify_def) lemma rett_instr_result: "(fst i) = ctrl_type RETT \ (get_ET (cpu_reg_val PSR s) \ 1 \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s)) = 0 \ ((AND) (get_addr (snd i) s) (0b00000000000000000000000000000011::word32)) = 0) \ snd (rett_instr i s) = False" apply (simp add: rett_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: raise_trap_def simpler_modify_def) by (simp add: return_def) lemma call_instr_result: "(fst i) = call_type CALL \ snd (call_instr i s) = False" apply (simp add: call_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def case_prod_unfold) apply (simp add: write_cpu_def write_reg_def) apply (simp add: get_curr_win_def get_CWP_def) by (simp add: simpler_modify_def simpler_gets_def) lemma branch_instr_result: "(fst i) \ {bicc_type BE,bicc_type BNE,bicc_type BGU, bicc_type BLE,bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA,bicc_type BN} \ snd (branch_instr i s) = False" proof (cases "eval_icc (fst i) (get_icc_N ((cpu_reg s) PSR)) (get_icc_Z ((cpu_reg s) PSR)) (get_icc_V ((cpu_reg s) PSR)) (get_icc_C ((cpu_reg s) PSR)) = 1") case True then have f1: "eval_icc (fst i) (get_icc_N ((cpu_reg s) PSR)) (get_icc_Z ((cpu_reg s) PSR)) (get_icc_V ((cpu_reg s) PSR)) (get_icc_C ((cpu_reg s) PSR)) = 1" by auto then show ?thesis proof (cases "(fst i) = bicc_type BA \ get_operand_flag ((snd i)!0) = 1") case True then show ?thesis using f1 apply (simp add: branch_instr_def) apply (simp add: Let_def simpler_gets_def bind_def h1_def h2_def) apply (simp add: set_annul_def case_prod_unfold) apply (simp add: write_cpu_def simpler_modify_def) by (simp add: return_def) next case False then have f2: "\ (fst i = bicc_type BA \ get_operand_flag (snd i ! 0) = 1)" by auto then show ?thesis using f1 apply (simp add: branch_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: branch_instr_sub1_def) apply (simp add: Let_def) apply auto apply (simp add: write_cpu_def simpler_modify_def) by (simp add: write_cpu_def simpler_modify_def) qed next case False then show ?thesis apply (simp add: branch_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: branch_instr_sub1_def) apply (simp add: Let_def) apply auto apply (simp add: Let_def bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: cpu_reg_mod_def set_annul_def simpler_modify_def) by (simp add: write_cpu_def simpler_modify_def) qed lemma nop_instr_result: "(fst i) = nop_type NOP \ snd (nop_instr i s) = False" apply (simp add: nop_instr_def) by (simp add: returnOk_def return_def) lemma sethi_instr_result: "(fst i) = sethi_type SETHI \ snd (sethi_instr i s) = False" apply (simp add: sethi_instr_def) apply (simp add: Let_def) apply (simp add: get_curr_win_def get_CWP_def cpu_reg_val_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: write_reg_def simpler_modify_def) by (simp add: return_def) lemma jmpl_instr_result: "(fst i) = ctrl_type JMPL \ snd (jmpl_instr i s) = False" apply (simp add: jmpl_instr_def) apply (simp add: get_curr_win_def get_CWP_def cpu_reg_val_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: write_cpu_def simpler_modify_def) by (simp add: raise_trap_def simpler_modify_def) lemma save_restore_instr_result: "(fst i) \ {ctrl_type SAVE,ctrl_type RESTORE} \ snd (save_restore_instr i s) = False" proof (cases "(fst i) = ctrl_type SAVE") case True then show ?thesis by save_restore_proof next case False then show ?thesis by save_restore_proof qed lemma flush_instr_result: "(fst i) = load_store_type FLUSH \ snd (flush_instr i s) = False" apply (simp add: flush_instr_def) by (simp add: simpler_gets_def bind_def h1_def h2_def simpler_modify_def) lemma read_state_reg_instr_result: "(fst i) \ {sreg_type RDY,sreg_type RDPSR, sreg_type RDWIM,sreg_type RDTBR} \ snd (read_state_reg_instr i s) = False" apply (simp add: read_state_reg_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_gets_def bind_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: raise_trap_def simpler_modify_def return_def) apply (simp add: bind_def h1_def h2_def) by (simp add: get_curr_win_def simpler_gets_def) lemma write_state_reg_instr_result: "(fst i) \ {sreg_type WRY,sreg_type WRPSR, sreg_type WRWIM,sreg_type WRTBR} \ snd (write_state_reg_instr i s) = False" apply (simp add: write_state_reg_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_modify_def) apply (simp add: raise_trap_def simpler_modify_def return_def) apply (simp add: bind_def h1_def h2_def) apply (simp add: simpler_gets_def) apply (simp add: write_cpu_def simpler_modify_def) by (simp add: get_curr_win_def simpler_gets_def) lemma logical_instr_result: "(fst i) \ {logic_type ANDs,logic_type ANDcc, logic_type ANDN,logic_type ANDNcc,logic_type ORs,logic_type ORcc, logic_type ORN,logic_type XORs,logic_type XNOR} \ snd (logical_instr i s) = False" apply (simp add: logical_instr_def) apply (simp add: Let_def simpler_gets_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: logical_instr_sub1_def) apply (simp add: return_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_gets_def) by (simp add: get_curr_win_def simpler_gets_def) lemma shift_instr_result: "(fst i) \ {shift_type SLL,shift_type SRL,shift_type SRA} \ snd (shift_instr i s) = False" apply (simp add: shift_instr_def) apply (simp add: Let_def) apply (simp add: get_curr_win_def simpler_gets_def bind_def h1_def h2_def) apply (simp add: return_def) apply (simp add: bind_def h1_def h2_def) by (simp add: write_reg_def simpler_modify_def) method add_sub_instr_proof = ((simp add: Let_def), auto, (simp add: write_reg_def simpler_modify_def), (simp add: simpler_gets_def bind_def), (simp add: get_curr_win_def simpler_gets_def), (simp add: write_reg_def write_cpu_def simpler_modify_def), (simp add: bind_def), (simp add: case_prod_unfold), (simp add: simpler_gets_def), (simp add: get_curr_win_def simpler_gets_def), (simp add: write_reg_def simpler_modify_def), (simp add: simpler_gets_def bind_def), (simp add: get_curr_win_def simpler_gets_def)) lemma add_instr_result: "(fst i) \ {arith_type ADD,arith_type ADDcc,arith_type ADDX} \ snd (add_instr i s) = False" apply (simp add: add_instr_def) apply (simp add: Let_def) apply auto apply (simp add: add_instr_sub1_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_gets_def) apply (simp add: get_curr_win_def simpler_gets_def) apply (simp add: add_instr_sub1_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: get_curr_win_def simpler_gets_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: add_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: get_curr_win_def simpler_gets_def) by (simp add: write_reg_def simpler_modify_def) lemma sub_instr_result: "(fst i) \ {arith_type SUB,arith_type SUBcc, arith_type SUBX} \ snd (sub_instr i s) = False" apply (simp add: sub_instr_def) apply (simp add: Let_def) apply auto apply (simp add: sub_instr_sub1_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_gets_def) apply (simp add: get_curr_win_def simpler_gets_def) apply (simp add: sub_instr_sub1_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: get_curr_win_def simpler_gets_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: sub_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: get_curr_win_def simpler_gets_def) by (simp add: write_reg_def simpler_modify_def) lemma mul_instr_result: "(fst i) \ {arith_type UMUL,arith_type SMUL, arith_type SMULcc} \ snd (mul_instr i s) = False" apply (simp add: mul_instr_def) apply (simp add: Let_def) apply auto apply (simp add: mul_instr_sub1_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: get_curr_win_def simpler_gets_def) apply (simp add: write_reg_def write_cpu_def simpler_modify_def) apply (simp add: mul_instr_sub1_def) apply (simp add: simpler_gets_def) apply (simp add: write_cpu_def write_reg_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: get_curr_win_def simpler_gets_def) apply (simp add: mul_instr_sub1_def) apply (simp add: simpler_gets_def) apply (simp add: write_cpu_def write_reg_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def) by (simp add: get_curr_win_def simpler_gets_def) lemma div_write_new_val_result: "snd (div_write_new_val i result temp_V s) = False" apply (simp add: div_write_new_val_def) apply (simp add: return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) by (simp add: write_cpu_def simpler_modify_def) lemma div_result: "snd (div_comp instr rs1 rd operand2 s) = False" apply (simp add: div_comp_def) apply (simp add: simpler_gets_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: get_curr_win_def simpler_gets_def) by (simp add: div_write_new_val_result) lemma div_instr_result: "(fst i) \ {arith_type UDIV,arith_type UDIVcc, arith_type SDIV} \ snd (div_instr i s) = False" apply (simp add: div_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: raise_trap_def simpler_modify_def) apply (simp add: return_def bind_def) by (simp add: div_result) lemma load_sub2_result: "snd (load_sub2 address asi rd curr_win word0 s) = False" apply (simp add: load_sub2_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def) apply (simp add: write_reg_def simpler_modify_def) by (simp add: simpler_gets_def) lemma load_sub3_result: "snd (load_sub3 instr curr_win rd asi address s) = False" apply (simp add: load_sub3_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: load_sub2_result) by (simp add: raise_trap_def simpler_modify_def) lemma load_sub1_result: "snd (load_sub1 i rd s_val s) = False" apply (simp add: load_sub1_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def simpler_modify_def) apply (simp add: get_curr_win_def simpler_gets_def) by (simp add: load_sub3_result) lemma load_instr_result: "(fst i) \ {load_store_type LDSB,load_store_type LDUB, load_store_type LDUBA,load_store_type LDUH,load_store_type LD, load_store_type LDA,load_store_type LDD} \ snd (load_instr i s) = False" apply (simp add: load_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: raise_trap_def simpler_modify_def) apply (simp add: return_def) by (simp add: load_sub1_result) lemma store_sub2_result: "snd (store_sub2 instr curr_win rd asi address s) = False" apply (simp add: store_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: raise_trap_def simpler_modify_def) apply (simp add: return_def) apply (simp add: raise_trap_def simpler_modify_def) by (simp add: bind_def h1_def h2_def) lemma store_sub1_result: "snd (store_sub1 instr rd s_val s) = False" apply (simp add: store_sub1_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def simpler_modify_def) apply (simp add: get_curr_win_def) apply (simp add: simpler_gets_def) by (simp add: store_sub2_result) lemma store_instr_result: "(fst i) \ {load_store_type STB,load_store_type STH, load_store_type ST,load_store_type STA,load_store_type STD} \ snd (store_instr i s) = False" apply (simp add: store_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: raise_trap_def simpler_modify_def) apply (simp add: return_def) by (simp add: store_sub1_result) lemma supported_instr_set: "supported_instruction i = True \ i \ {load_store_type LDSB,load_store_type LDUB,load_store_type LDUBA, load_store_type LDUH,load_store_type LD,load_store_type LDA, load_store_type LDD, load_store_type STB,load_store_type STH,load_store_type ST, load_store_type STA,load_store_type STD, sethi_type SETHI, nop_type NOP, logic_type ANDs,logic_type ANDcc,logic_type ANDN,logic_type ANDNcc, logic_type ORs,logic_type ORcc,logic_type ORN,logic_type XORs, logic_type XNOR, shift_type SLL,shift_type SRL,shift_type SRA, arith_type ADD,arith_type ADDcc,arith_type ADDX, arith_type SUB,arith_type SUBcc,arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, ctrl_type RETT, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" apply (simp add: supported_instruction_def) by presburger lemma dispatch_instr_result: assumes a1: "supported_instruction (fst i) = True \ (fst i) \ ctrl_type RETT" shows "snd (dispatch_instruction i s) = False" proof (cases "get_trap_set s = {}") case True then have f1: "get_trap_set s = {}" by auto then show ?thesis proof (cases "(fst i) \ {load_store_type LDSB,load_store_type LDUB, load_store_type LDUBA,load_store_type LDUH,load_store_type LD, load_store_type LDA,load_store_type LDD}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (simp add: load_instr_result) next case False then have f2: "(fst i) \ {load_store_type STB,load_store_type STH,load_store_type ST, load_store_type STA,load_store_type STD, sethi_type SETHI, nop_type NOP, logic_type ANDs,logic_type ANDcc,logic_type ANDN,logic_type ANDNcc, logic_type ORs,logic_type ORcc,logic_type ORN,logic_type XORs, logic_type XNOR, shift_type SLL,shift_type SRL,shift_type SRA, arith_type ADD,arith_type ADDcc,arith_type ADDX, arith_type SUB,arith_type SUBcc,arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using a1 apply (simp add: supported_instruction_def) by presburger then show ?thesis proof (cases "(fst i) \ {load_store_type STB,load_store_type STH, load_store_type ST, load_store_type STA,load_store_type STD}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: store_instr_result) next case False then have f3: "(fst i) \ {sethi_type SETHI, nop_type NOP, logic_type ANDs,logic_type ANDcc,logic_type ANDN,logic_type ANDNcc, logic_type ORs,logic_type ORcc,logic_type ORN,logic_type XORs, logic_type XNOR, shift_type SLL,shift_type SRL,shift_type SRA, arith_type ADD,arith_type ADDcc,arith_type ADDX, arith_type SUB,arith_type SUBcc,arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f2 by auto then show ?thesis proof (cases "(fst i) = sethi_type SETHI") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (simp add: sethi_instr_result) next case False then have f4: "(fst i) \ {nop_type NOP, logic_type ANDs,logic_type ANDcc,logic_type ANDN,logic_type ANDNcc, logic_type ORs,logic_type ORcc,logic_type ORN,logic_type XORs, logic_type XNOR, shift_type SLL,shift_type SRL,shift_type SRA, arith_type ADD,arith_type ADDcc,arith_type ADDX, arith_type SUB,arith_type SUBcc,arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f3 by auto then show ?thesis proof (cases "fst i = nop_type NOP") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (simp add: nop_instr_result) next case False then have f5: "(fst i) \ {logic_type ANDs,logic_type ANDcc, logic_type ANDN,logic_type ANDNcc, logic_type ORs,logic_type ORcc,logic_type ORN,logic_type XORs, logic_type XNOR, shift_type SLL,shift_type SRL,shift_type SRA, arith_type ADD,arith_type ADDcc,arith_type ADDX, arith_type SUB,arith_type SUBcc,arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f4 by auto then show ?thesis proof (cases "(fst i) \ {logic_type ANDs,logic_type ANDcc, logic_type ANDN,logic_type ANDNcc, logic_type ORs,logic_type ORcc,logic_type ORN,logic_type XORs, logic_type XNOR}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: logical_instr_result) next case False then have f6: "(fst i) \ {shift_type SLL,shift_type SRL, shift_type SRA, arith_type ADD,arith_type ADDcc,arith_type ADDX, arith_type SUB,arith_type SUBcc,arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f5 by auto then show ?thesis proof (cases "(fst i) \ {shift_type SLL,shift_type SRL, shift_type SRA}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: shift_instr_result) next case False then have f7: "(fst i) \ {arith_type ADD,arith_type ADDcc, arith_type ADDX, arith_type SUB,arith_type SUBcc,arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f6 by auto then show ?thesis proof (cases "(fst i) \ {arith_type ADD,arith_type ADDcc, arith_type ADDX}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: add_instr_result) next case False then have f8: "(fst i) \ {arith_type SUB,arith_type SUBcc, arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f7 by auto then show ?thesis proof (cases "(fst i) \ {arith_type SUB,arith_type SUBcc, arith_type SUBX}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: sub_instr_result) next case False then have f9: "(fst i) \ {arith_type UMUL,arith_type SMUL, arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f8 by auto then show ?thesis proof (cases "(fst i) \ {arith_type UMUL,arith_type SMUL, arith_type SMULcc}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: mul_instr_result) next case False then have f10: "(fst i) \ {arith_type UDIV,arith_type UDIVcc, arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f9 by auto then show ?thesis proof (cases "(fst i) \ {arith_type UDIV,arith_type UDIVcc, arith_type SDIV}") case True then show ?thesis apply dispatch_instr_proof1 using f1 by (auto simp add: div_instr_result) next case False then have f11: "(fst i) \ {ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f10 by auto then show ?thesis proof (cases "(fst i) \ {ctrl_type SAVE,ctrl_type RESTORE}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: save_restore_instr_result) next case False then have f12: "(fst i) \ {call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f11 by auto then show ?thesis proof (cases "(fst i) = call_type CALL") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: call_instr_result) next case False then have f13: "(fst i) \ {ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f12 by auto then show ?thesis proof (cases "(fst i) = ctrl_type JMPL") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: jmpl_instr_result) next case False then have f14: "(fst i) \ { sreg_type RDY, sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f13 by auto then show ?thesis proof (cases "(fst i) \ {sreg_type RDY, sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: read_state_reg_instr_result) next case False then have f15: "(fst i) \ { sreg_type WRY, sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f14 by auto then show ?thesis proof (cases "(fst i) \ {sreg_type WRY, sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: write_state_reg_instr_result) next case False then have f16: "(fst i) \ { load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f15 by auto then show ?thesis proof (cases "(fst i) = load_store_type FLUSH") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: flush_instr_result) next case False then have f17: "(fst i) \ { bicc_type BE,bicc_type BNE,bicc_type BGU, bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f16 by auto then show ?thesis using f1 proof (cases "(fst i) \ {bicc_type BE, bicc_type BNE,bicc_type BGU, bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA }") case True then show ?thesis using f1 apply dispatch_instr_proof1 apply auto by (auto simp add: branch_instr_result) next case False then have f18: "(fst i) \ {bicc_type BN}" using f17 by auto then show ?thesis using f1 apply dispatch_instr_proof1 apply auto by (auto simp add: branch_instr_result) qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed next case False then show ?thesis apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) by (simp add: returnOk_def return_def) qed lemma dispatch_instr_result_rett: assumes a1: "(fst i) = ctrl_type RETT \ (get_ET (cpu_reg_val PSR s) \ 1 \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s)) = 0 \ ((AND) (get_addr (snd i) s) (0b00000000000000000000000000000011::word32)) = 0)" shows "snd (dispatch_instruction i s) = False" proof (cases "get_trap_set s = {}") case True then show ?thesis using a1 apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: rett_instr_result) next case False then show ?thesis using a1 apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: return_def) qed lemma execute_instr_sub1_result: "snd (execute_instr_sub1 i s) = False" proof (cases "get_trap_set s = {} \ (fst i) \ {call_type CALL,ctrl_type RETT, ctrl_type JMPL}") case True then show ?thesis apply (simp add: execute_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) apply auto by (auto simp add: return_def) next case False then show ?thesis apply (simp add: execute_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) by (auto simp add: return_def) qed lemma next_match : "snd (execute_instruction () s) = False \ NEXT s = Some (snd (fst (execute_instruction () s)))" apply (simp add: NEXT_def) by (simp add: case_prod_unfold) lemma exec_ss1 : "\s'. (execute_instruction () s = (s', False)) \ \s''. (execute_instruction() s = (s'', False))" proof - assume "\s'. (execute_instruction () s = (s', False))" hence "(snd (execute_instruction() s)) = False" by (auto simp add: execute_instruction_def case_prod_unfold) hence "(execute_instruction() s) = ((fst (execute_instruction() s)),False)" by (metis (full_types) prod.collapse) hence "\s''. (execute_instruction() s = (s'', False))" by blast thus ?thesis by assumption qed lemma exec_ss2 : "snd (execute_instruction() s) = False \ snd (execute_instruction () s) = False" proof - assume "snd (execute_instruction() s) = False" hence "snd (execute_instruction () s) = False" by (auto simp add:execute_instruction_def) thus ?thesis by assumption qed lemma good_context_1 : "good_context s \ s' = s \ (get_trap_set s') \ {} \ (reset_trap_val s') = False \ get_ET (cpu_reg_val PSR s') = 0 \ False" proof - assume asm: "good_context s \ s' = s \ (get_trap_set s') \ {} \ (reset_trap_val s') = False \ get_ET (cpu_reg_val PSR s') = 0" then have "(get_trap_set s') \ {} \ (reset_trap_val s') = False \ get_ET (cpu_reg_val PSR s') = 0 \ False" by (simp add: good_context_def get_ET_def cpu_reg_val_def) then show ?thesis using asm by auto qed lemma fetch_instr_result_1 : "\ (\e. fetch_instruction s' = Inl e) \ (\v. fetch_instruction s' = Inr v)" by (meson sumE) lemma fetch_instr_result_2 : "(\v. fetch_instruction s' = Inr v) \ \ (\e. fetch_instruction s' = Inl e)" by force lemma fetch_instr_result_3 : "(\e. fetch_instruction s' = Inl e) \ \ (\v. fetch_instruction s' = Inr v)" by auto lemma decode_instr_result_1 : "\(\v2. ((decode_instruction v1)::(Exception list + instruction)) = Inr v2) \ (\e. ((decode_instruction v1)::(Exception list + instruction)) = Inl e)" by (meson sumE) lemma decode_instr_result_2 : "(\e. ((decode_instruction v1)::(Exception list + instruction)) = Inl e) \ \(\v2. ((decode_instruction v1)::(Exception list + instruction)) = Inr v2)" by force lemma decode_instr_result_3 : "x = decode_instruction v1 \ y = decode_instruction v2 \ v1 = v2 \ x = y" by auto lemma decode_instr_result_4 : "\ (\e. ((decode_instruction v1)::(Exception list + instruction)) = Inl e) \ (\v2. ((decode_instruction v1)::(Exception list + instruction)) = Inr v2)" by (meson sumE) lemma good_context_2 : -"good_context (s::(('a::len0) sparc_state)) \ +"good_context (s::(('a::len) sparc_state)) \ fetch_instruction (delayed_pool_write s) = Inr v1 \ \(\v2. (decode_instruction v1::(Exception list + instruction)) = Inr v2) \ False" proof - assume "good_context s \ fetch_instruction (delayed_pool_write s) = Inr v1 \ \(\v2. ((decode_instruction v1)::(Exception list + instruction)) = Inr v2)" hence fact1: "good_context s \ fetch_instruction (delayed_pool_write s) = Inr v1 \ (\e. ((decode_instruction v1)::(Exception list + instruction)) = Inl e)" using decode_instr_result_1 by auto hence fact2: "\(\e. fetch_instruction (delayed_pool_write s) = Inl e)" using fetch_instr_result_2 by auto then have "fetch_instruction (delayed_pool_write s) = Inr v1 \ (\e. ((decode_instruction v1)::(Exception list + instruction)) = Inl e) \ False" proof (cases "(get_trap_set s) \ {} \ (reset_trap_val s) = False \ get_ET (cpu_reg_val PSR s) = 0") case True from this fact1 show ?thesis using good_context_1 by blast next case False then have fact3: "(get_trap_set s) = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0" by auto then show ?thesis using fact1 decode_instr_result_3 by (metis (no_types, lifting) good_context_def sum.case(1) sum.case(2)) qed thus ?thesis using fact1 by auto qed lemma good_context_3 : -"good_context (s::(('a::len0) sparc_state)) \ +"good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ (decode_instruction v1::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = False \ False" proof - - assume asm: "good_context (s::(('a::len0) sparc_state)) \ + assume asm: "good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ (decode_instruction v1::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = False" then have "annul_val s'' = False \ supported_instruction (fst v2) = False \ False" proof (cases "(get_trap_set s) \ {} \ (reset_trap_val s) = False \ get_ET (cpu_reg_val PSR s) = 0") case True from this asm show ?thesis using good_context_1 by blast next case False then have fact3: "(get_trap_set s) = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0" by auto thus ?thesis using asm by (auto simp add: good_context_def) qed thus ?thesis using asm by auto qed lemma good_context_4 : -"good_context (s::(('a::len0) sparc_state)) \ +"good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ \ \This line is redundant\ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ ((ucast (get_S (cpu_reg_val PSR s'')))::word1) = 0 \ False" proof - - assume asm: "good_context (s::(('a::len0) sparc_state)) \ + assume asm: "good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ \ \This line is redundant\ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ ((ucast (get_S (cpu_reg_val PSR s'')))::word1) = 0" then have "(fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ ((ucast (get_S (cpu_reg_val PSR s'')))::word1) = 0 \ False" proof (cases "(get_trap_set s) \ {} \ (reset_trap_val s) = False \ get_ET (cpu_reg_val PSR s) = 0") case True from this asm show ?thesis using good_context_1 by blast next case False then have fact3: "(get_trap_set s) = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0" by auto thus ?thesis using asm by (auto simp add: good_context_def) qed thus ?thesis using asm by auto qed lemma good_context_5 : -"good_context (s::(('a::len0) sparc_state)) \ +"good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ \ \This line is redundant\ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ ((ucast (get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) \ 0 \ False" proof - - assume asm: "good_context (s::(('a::len0) sparc_state)) \ + assume asm: "good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ \ \This line is redundant\ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ ((ucast (get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) \ 0" then have "(fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ ((ucast (get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) \ 0 \ False" proof (cases "(get_trap_set s) \ {} \ (reset_trap_val s) = False \ get_ET (cpu_reg_val PSR s) = 0") case True from this asm show ?thesis using good_context_1 by blast next case False then have fact3: "(get_trap_set s) = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0" by auto thus ?thesis using asm by (auto simp add: good_context_def) qed thus ?thesis using asm by auto qed lemma good_context_6 : -"good_context (s::(('a::len0) sparc_state)) \ +"good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ \ \This line is redundant\ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ ((ucast (get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) = 0 \ ((AND) (get_addr (snd v2) s'') (0b00000000000000000000000000000011::word32)) \ 0 \ False" proof - - assume asm: "good_context (s::(('a::len0) sparc_state)) \ + assume asm: "good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ \ \This line is redundant\ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ ((ucast (get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) = 0 \ ((AND) (get_addr (snd v2) s'') (0b00000000000000000000000000000011::word32)) \ 0" then have "(fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ ((ucast (get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) = 0 \ ((AND) (get_addr (snd v2) s'') (0b00000000000000000000000000000011::word32)) \ 0 \ False" proof (cases "(get_trap_set s) \ {} \ (reset_trap_val s) = False \ get_ET (cpu_reg_val PSR s) = 0") case True from this asm show ?thesis using good_context_1 by blast next case False then have fact3: "(get_trap_set s) = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0" by auto thus ?thesis using asm by (auto simp add: good_context_def) qed thus ?thesis using asm by auto qed lemma good_context_all : -"good_context (s::(('a::len0) sparc_state)) \ +"good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ (get_trap_set s = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0) \ ((\e. fetch_instruction s'' = Inl e) \ (\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (annul_val s'' = True \ (annul_val s'' = False \ (\v1' v2'. fetch_instruction s'' = Inr v1' \ ((decode_instruction v1')::(Exception list + instruction)) = Inr v2' \ supported_instruction (fst v2') = True) \ ((fst v2) \ ctrl_type RETT \ ((fst v2) = ctrl_type RETT \ (get_ET (cpu_reg_val PSR s'') = 1 \ (get_ET (cpu_reg_val PSR s'') \ 1 \ ((ucast (get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) = 0 \ ((AND) (get_addr (snd v2) s'') (0b00000000000000000000000000000011::word32)) = 0))))))))" proof - assume asm: "good_context s \ s'' = delayed_pool_write s" from asm have "(get_trap_set s) \ {} \ (reset_trap_val s) = False \ get_ET (cpu_reg_val PSR s) = 0 \ False" using good_context_1 by blast hence fact1: "(get_trap_set s = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0)" by auto have fact2: "\(\e. fetch_instruction s'' = Inl e) \ \ (\v1. fetch_instruction s'' = Inr v1) \ False" using fetch_instr_result_1 by blast from asm have fact3: "\v1. fetch_instruction s'' = Inr v1 \ \(\v2.((decode_instruction v1)::(Exception list + instruction)) = Inr v2) \ False" using good_context_2 by blast from asm have fact4: "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = False \ False" using good_context_3 by blast from asm have fact5: "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ ((ucast (get_S (cpu_reg_val PSR s'')))::word1) = 0 \ False" using good_context_4 by blast from asm have fact6: "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ ((ucast (get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) \ 0 \ False" using good_context_5 by blast from asm have fact7: "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ ((ucast (get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) = 0 \ ((AND) (get_addr (snd v2) s'') (0b00000000000000000000000000000011::word32)) \ 0 \ False" using good_context_6 by blast from asm show ?thesis proof (cases "(\e. fetch_instruction s'' = Inl e)") case True then show ?thesis using fact1 by auto next case False then have fact8: "\v1. fetch_instruction s'' = Inr v1 \ (\v2.((decode_instruction v1)::(Exception list + instruction)) = Inr v2)" using fact2 fact3 by auto then show ?thesis proof (cases "annul_val s'' = True") case True then show ?thesis using fact1 fact8 by auto next case False then have fact9: "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True" using fact4 fact8 by blast then show ?thesis proof (cases "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (fst v2) \ ctrl_type RETT") case True then show ?thesis using fact1 fact9 by auto next case False then have fact10: "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ (fst v2) = ctrl_type RETT" using fact9 by auto then show ?thesis proof (cases "get_ET (cpu_reg_val PSR s'') = 1") case True then show ?thesis using fact1 fact9 by auto next case False then have fact11: "get_ET (cpu_reg_val PSR s'') \ 1 \ ((ucast (get_S (cpu_reg_val PSR s'')))::word1) \ 0" using fact10 fact5 by auto then have fact12: "(get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) = 0" using fact10 fact6 by auto then have fact13: "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ ((AND) (get_addr (snd v2) s'') (0b00000000000000000000000000000011::word32)) = 0" using fact10 fact11 fact7 by blast thus ?thesis using fact1 fact10 fact11 fact12 by auto qed qed qed qed qed lemma select_trap_result1 : "(reset_trap_val s) = True \ snd (select_trap() s) = False" apply (simp add: select_trap_def exec_gets return_def) by (simp add: bind_def h1_def h2_def simpler_modify_def) lemma select_trap_result2 : assumes a1: "\(reset_trap_val s = False \ get_ET (cpu_reg_val PSR s) = 0)" shows "snd (select_trap() s) = False" proof (cases "reset_trap_val s = True") case True then show ?thesis using select_trap_result1 by blast next case False then have f1: "reset_trap_val s = False \ get_ET (cpu_reg_val PSR s) \ 0" using a1 by auto then show ?thesis proof (cases "data_store_error \ get_trap_set s") case True then show ?thesis using f1 by select_trap_proof0 next case False then have f2: "data_store_error \ get_trap_set s" by auto then show ?thesis proof (cases "instruction_access_error \ get_trap_set s") case True then show ?thesis using f1 f2 by select_trap_proof0 next case False then have f3: "instruction_access_error \ get_trap_set s" by auto then show ?thesis proof (cases "r_register_access_error \ get_trap_set s") case True then show ?thesis using f1 f2 f3 by select_trap_proof0 next case False then have f4: "r_register_access_error \ get_trap_set s" by auto then show ?thesis proof (cases "instruction_access_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 by select_trap_proof0 next case False then have f5: "instruction_access_exception \ get_trap_set s" by auto then show ?thesis proof (cases "privileged_instruction \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 by select_trap_proof0 next case False then have f6: "privileged_instruction \ get_trap_set s" by auto then show ?thesis proof (cases "illegal_instruction \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 by select_trap_proof0 next case False then have f7: "illegal_instruction \ get_trap_set s" by auto then show ?thesis proof (cases "fp_disabled \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 by select_trap_proof0 next case False then have f8: "fp_disabled \ get_trap_set s" by auto then show ?thesis proof (cases "cp_disabled \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 by select_trap_proof0 next case False then have f9: "cp_disabled \ get_trap_set s" by auto then show ?thesis proof (cases "unimplemented_FLUSH \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 by select_trap_proof0 next case False then have f10: "unimplemented_FLUSH \ get_trap_set s" by auto then show ?thesis proof (cases "window_overflow \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 by select_trap_proof0 next case False then have f11: "window_overflow \ get_trap_set s" by auto then show ?thesis proof (cases "window_underflow \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 by select_trap_proof0 next case False then have f12: "window_underflow \ get_trap_set s" by auto then show ?thesis proof (cases "mem_address_not_aligned \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 by select_trap_proof0 next case False then have f13: "mem_address_not_aligned \ get_trap_set s" by auto then show ?thesis proof (cases "fp_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 by select_trap_proof0 next case False then have f14: "fp_exception \ get_trap_set s" by auto then show ?thesis proof (cases "cp_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 by select_trap_proof0 next case False then have f15: "cp_exception \ get_trap_set s" by auto then show ?thesis proof (cases "data_access_error \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 by select_trap_proof0 next case False then have f16: "data_access_error \ get_trap_set s" by auto then show ?thesis proof (cases "data_access_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 by select_trap_proof0 next case False then have f17: "data_access_exception \ get_trap_set s" by auto then show ?thesis proof (cases "tag_overflow \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 by select_trap_proof0 next case False then have f18: "tag_overflow \ get_trap_set s" by auto then show ?thesis proof (cases "division_by_zero \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 by select_trap_proof0 next case False then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 apply (simp add: select_trap_def exec_gets return_def) apply (simp add: DetMonad.bind_def h1_def h2_def simpler_modify_def) apply (simp add: return_def simpler_gets_def) apply (simp add: case_prod_unfold) apply (simp add: return_def) apply (simp add: write_cpu_tt_def write_cpu_def) by (simp add: simpler_gets_def bind_def h1_def h2_def simpler_modify_def) qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed lemma emp_trap_set_err_mode : "err_mode_val s = err_mode_val (emp_trap_set s)" by (auto simp add: emp_trap_set_def err_mode_val_def) lemma write_cpu_tt_err_mode : "err_mode_val s = err_mode_val (snd (fst (write_cpu_tt w s)))" apply (simp add: write_cpu_tt_def err_mode_val_def write_cpu_def) apply (simp add: exec_gets return_def) apply (simp add: bind_def simpler_modify_def) by (simp add: cpu_reg_mod_def) lemma select_trap_monad : "snd (select_trap() s) = False \ err_mode_val s = err_mode_val (snd (fst (select_trap () s)))" proof - assume a1: "snd (select_trap() s) = False" then have f0: "reset_trap_val s = False \ get_ET (cpu_reg_val PSR s) = 0 \ False" apply (simp add: select_trap_def exec_gets return_def) apply (simp add: bind_def h1_def h2_def simpler_modify_def) by (simp add: fail_def split_def) then show ?thesis proof (cases "reset_trap_val s = True") case True from a1 f0 this show ?thesis apply (simp add: select_trap_def exec_gets return_def) apply (simp add: bind_def h1_def h2_def simpler_modify_def) by (simp add: emp_trap_set_def err_mode_val_def) next case False then have f1: "reset_trap_val s = False \ get_ET (cpu_reg_val PSR s) \ 0" using f0 by auto then show ?thesis using f1 a1 proof (cases "data_store_error \ get_trap_set s") case True then show ?thesis using f1 a1 by select_trap_proof1 next case False then have f2: "data_store_error \ get_trap_set s" by auto then show ?thesis proof (cases "instruction_access_error \ get_trap_set s") case True then show ?thesis using f1 f2 a1 by select_trap_proof1 next case False then have f3: "instruction_access_error \ get_trap_set s" by auto then show ?thesis proof (cases "r_register_access_error \ get_trap_set s") case True then show ?thesis using f1 f2 f3 a1 by select_trap_proof1 next case False then have f4: "r_register_access_error \ get_trap_set s" by auto then show ?thesis proof (cases "instruction_access_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 a1 by select_trap_proof1 next case False then have f5: "instruction_access_exception \ get_trap_set s" by auto then show ?thesis proof (cases "privileged_instruction \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 a1 by select_trap_proof1 next case False then have f6: "privileged_instruction \ get_trap_set s" by auto then show ?thesis proof (cases "illegal_instruction \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 a1 by select_trap_proof1 next case False then have f7: "illegal_instruction \ get_trap_set s" by auto then show ?thesis proof (cases "fp_disabled \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 a1 by select_trap_proof1 next case False then have f8: "fp_disabled \ get_trap_set s" by auto then show ?thesis proof (cases "cp_disabled \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 a1 by select_trap_proof1 next case False then have f9: "cp_disabled \ get_trap_set s" by auto then show ?thesis proof (cases "unimplemented_FLUSH \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 a1 by select_trap_proof1 next case False then have f10: "unimplemented_FLUSH \ get_trap_set s" by auto then show ?thesis proof (cases "window_overflow \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 a1 by select_trap_proof1 next case False then have f11: "window_overflow \ get_trap_set s" by auto then show ?thesis proof (cases "window_underflow \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 a1 by select_trap_proof1 next case False then have f12: "window_underflow \ get_trap_set s" by auto then show ?thesis proof (cases "mem_address_not_aligned \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 a1 by select_trap_proof1 next case False then have f13: "mem_address_not_aligned \ get_trap_set s" by auto then show ?thesis proof (cases "fp_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 a1 by select_trap_proof1 next case False then have f14: "fp_exception \ get_trap_set s" by auto then show ?thesis proof (cases "cp_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 a1 by select_trap_proof1 next case False then have f15: "cp_exception \ get_trap_set s" by auto then show ?thesis proof (cases "data_access_error \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 a1 by select_trap_proof1 next case False then have f16: "data_access_error \ get_trap_set s" by auto then show ?thesis proof (cases "data_access_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 a1 by select_trap_proof1 next case False then have f17: "data_access_exception \ get_trap_set s" by auto then show ?thesis proof (cases "tag_overflow \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 a1 by select_trap_proof1 next case False then have f18: "tag_overflow \ get_trap_set s" by auto then show ?thesis proof (cases "division_by_zero \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 a1 by select_trap_proof1 next case False then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 a1 apply (simp add: select_trap_def exec_gets return_def) apply (simp add: bind_def h1_def h2_def simpler_modify_def) apply (simp add: return_def simpler_gets_def) apply (simp add: emp_trap_set_def err_mode_val_def cpu_reg_mod_def) apply (simp add: case_prod_unfold) apply (simp add: return_def) apply clarsimp apply (simp add: write_cpu_tt_def write_cpu_def write_tt_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: simpler_modify_def) by (simp add: cpu_reg_val_def cpu_reg_mod_def) qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed lemma exe_trap_st_pc_result : "snd (exe_trap_st_pc() s) = False" proof (cases "annul_val s = True") case True then show ?thesis apply (simp add: exe_trap_st_pc_def get_curr_win_def) apply (simp add: exec_gets return_def) apply (simp add: DetMonad.bind_def h1_def h2_def) by (simp add: set_annul_def write_reg_def simpler_modify_def) next case False then show ?thesis apply (simp add: exe_trap_st_pc_def get_curr_win_def) apply (simp add: exec_gets return_def) apply (simp add: DetMonad.bind_def h1_def h2_def) by (simp add: write_reg_def simpler_modify_def) qed lemma exe_trap_wr_pc_result : "snd (exe_trap_wr_pc() s) = False" proof (cases "reset_trap_val s = True") case True then show ?thesis apply (simp add: exe_trap_wr_pc_def get_curr_win_def) apply (simp add: exec_gets return_def) apply (simp add: DetMonad.bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: simpler_gets_def) apply (simp add: cpu_reg_val_def update_S_def cpu_reg_mod_def reset_trap_val_def) apply (simp add: write_cpu_def simpler_modify_def DetMonad.bind_def h1_def h2_def) apply (simp add: return_def) by (simp add: set_reset_trap_def simpler_modify_def DetMonad.bind_def h1_def h2_def return_def) next case False then show ?thesis apply (simp add: exe_trap_wr_pc_def get_curr_win_def) apply (simp add: exec_gets return_def) apply (simp add: DetMonad.bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: simpler_gets_def) apply (simp add: cpu_reg_val_def update_S_def cpu_reg_mod_def reset_trap_val_def) apply (simp add: write_cpu_def simpler_modify_def DetMonad.bind_def h1_def h2_def) by (simp add: return_def) qed lemma execute_trap_result : "\(reset_trap_val s = False \ get_ET (cpu_reg_val PSR s) = 0) \ snd (execute_trap() s) = False" proof - assume "\(reset_trap_val s = False \ get_ET (cpu_reg_val PSR s) = 0)" then have fact1: "snd (select_trap() s) = False" using select_trap_result2 by blast then show ?thesis proof (cases "err_mode_val s = True") case True then show ?thesis using fact1 apply (simp add: execute_trap_def exec_gets return_def) apply (simp add: DetMonad.bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (simp add: in_gets return_def select_trap_monad simpler_gets_def) next case False then show ?thesis using fact1 select_trap_monad apply (simp add: execute_trap_def exec_gets return_def) apply (simp add: DetMonad.bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_gets_def) apply (auto simp add: select_trap_monad) apply (simp add: DetMonad.bind_def h1_def h2_def get_curr_win_def) apply (simp add: get_CWP_def cpu_reg_val_def) apply (simp add: simpler_gets_def return_def write_cpu_def) apply (simp add: simpler_modify_def DetMonad.bind_def h1_def h2_def) apply (simp add: exe_trap_st_pc_result) by (simp add: case_prod_unfold exe_trap_wr_pc_result) qed qed lemma execute_trap_result2 : "\(reset_trap_val s = False \ get_ET (cpu_reg_val PSR s) = 0) \ snd (execute_trap() s) = False" using execute_trap_result by blast lemma exe_instr_all : -"good_context (s::(('a::len0) sparc_state)) \ +"good_context (s::(('a::len) sparc_state)) \ snd (execute_instruction() s) = False" proof - assume asm1: "good_context s" let ?s' = "delayed_pool_write s" from asm1 have f1 : "(get_trap_set s = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0) \ ((\e. fetch_instruction ?s' = Inl e) \ (\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (annul_val ?s' = True \ (annul_val ?s' = False \ (\v1' v2'. fetch_instruction ?s' = Inr v1' \ ((decode_instruction v1')::(Exception list + instruction)) = Inr v2' \ supported_instruction (fst v2') = True) \ ((fst v2) \ ctrl_type RETT \ ((fst v2) = ctrl_type RETT \ (get_ET (cpu_reg_val PSR ?s') = 1 \ (get_ET (cpu_reg_val PSR ?s') \ 1 \ ((ucast (get_S (cpu_reg_val PSR ?s')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR ?s'))) + 1) mod NWINDOWS)) (cpu_reg_val WIM ?s')) = 0 \ ((AND) (get_addr (snd v2) ?s') (0b00000000000000000000000000000011::word32)) = 0))))))))" using good_context_all by blast from f1 have f2: "get_trap_set s \ {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0" by auto show ?thesis proof (cases "get_trap_set s = {}") case True then have f3: "get_trap_set s = {}" by auto then show ?thesis proof (cases "exe_mode_val s = True") case True then have f4: "exe_mode_val s = True" by auto then show ?thesis proof (cases "\e1. fetch_instruction ?s' = Inl e1") case True then show ?thesis using f3 apply exe_proof_to_decode apply (simp add: raise_trap_def simpler_modify_def) by (simp add: bind_def h1_def h2_def return_def) next case False then have f5: "\ v1. fetch_instruction ?s' = Inr v1" using fetch_instr_result_1 by blast then have f6: "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2" using f1 fetch_instr_result_2 by blast then show ?thesis proof (cases "annul_val ?s' = True") case True then show ?thesis using f3 f4 f6 apply exe_proof_to_decode apply (simp add: set_annul_def annul_mod_def simpler_modify_def bind_def h1_def h2_def) apply (simp add: return_def simpler_gets_def) by (simp add: write_cpu_def simpler_modify_def) next case False then have f7: "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (\v1' v2'. fetch_instruction ?s' = Inr v1' \ ((decode_instruction v1')::(Exception list + instruction)) = Inr v2' \ supported_instruction (fst v2') = True) \ annul_val ?s' = False" using f1 f6 fetch_instr_result_2 by auto then have f7': "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ supported_instruction (fst v2) = True \ annul_val ?s' = False" by auto then show ?thesis proof (cases "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (fst v2) = ctrl_type RETT") case True then have f8: "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (fst v2) = ctrl_type RETT" by auto then show ?thesis proof (cases "get_trap_set ?s' = {}") case True then have f9: "get_trap_set ?s' = {}" by auto then show ?thesis proof (cases "get_ET (cpu_reg_val PSR ?s') = 1") case True then have f10: "get_ET (cpu_reg_val PSR ?s') = 1" by auto then show ?thesis proof (cases "((ucast (get_S (cpu_reg_val PSR ?s')))::word1) = 0") case True then show ?thesis using f3 f4 f7 f8 f9 f10 apply exe_proof_to_decode apply exe_proof_dispatch_rett apply (simp add: raise_trap_def simpler_modify_def) apply (auto simp add: execute_instr_sub1_result return_def) by (simp add: case_prod_unfold) next case False then show ?thesis using f3 f4 f7 f8 f9 f10 apply exe_proof_to_decode apply exe_proof_dispatch_rett apply (simp add: raise_trap_def simpler_modify_def) apply (auto simp add: execute_instr_sub1_result return_def) by (simp add: case_prod_unfold) qed next case False then have f11: "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val ?s' = False \ (fst v2) = ctrl_type RETT \ (get_ET (cpu_reg_val PSR ?s') \ 1 \ ((ucast (get_S (cpu_reg_val PSR ?s')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR ?s'))) + 1) mod NWINDOWS)) (cpu_reg_val WIM ?s')) = 0 \ ((AND) (get_addr (snd v2) ?s') (0b00000000000000000000000000000011::word32)) = 0)" using f1 fetch_instr_result_2 f7' f8 by auto then show ?thesis using f3 f4 proof (cases "get_trap_set ?s' = {}") case True then show ?thesis using f3 f4 f11 apply (simp add: execute_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def simpler_modify_def) apply clarsimp apply (simp add: return_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply auto apply (simp add: execute_instr_sub1_result) apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: rett_instr_result) next case False then show ?thesis using f3 f4 f11 apply (simp add: execute_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def simpler_modify_def) apply clarsimp apply (simp add: return_def) apply (simp add: bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply (simp add: execute_instr_sub1_result) apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: return_def) qed qed next case False then show ?thesis using f3 f4 f7 f8 apply exe_proof_to_decode apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) by (auto simp add: execute_instr_sub1_result return_def Let_def) qed next case False \ \Instruction is not \RETT\.\ then have "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (fst v2) \ ctrl_type RETT" using f7 by auto then have "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (fst v2) \ ctrl_type RETT \ supported_instruction (fst v2) = True \ annul_val ?s' = False" using f7 by auto then have "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (fst v2) \ ctrl_type RETT \ supported_instruction (fst v2) = True \ annul_val ?s' = False \ snd (dispatch_instruction v2 ?s') = False" by (auto simp add: dispatch_instr_result) then show ?thesis using f3 f4 apply exe_proof_to_decode apply (simp add: bind_def h1_def h2_def) apply (simp add: case_prod_unfold) by (simp add: execute_instr_sub1_result) qed qed qed next case False then show ?thesis using f3 apply (simp add: execute_instruction_def) by (simp add: exec_gets return_def) qed next case False then have "get_trap_set s \ {} \ ((reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0)" using f2 by auto then show ?thesis apply (simp add: execute_instruction_def exec_gets) by (simp add: execute_trap_result2) qed qed lemma dispatch_fail: -"snd (execute_instruction() (s::(('a::len0) sparc_state))) = False \ +"snd (execute_instruction() (s::(('a::len) sparc_state))) = False \ get_trap_set s = {} \ exe_mode_val s \ fetch_instruction (delayed_pool_write s) = Inr v \ ((decode_instruction v)::(Exception list + instruction)) = Inl e \ False" using decode_instr_result_2 apply (simp add: execute_instruction_def) apply (simp add: exec_gets bind_def) apply clarsimp apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: simpler_modify_def return_def) by (simp add: fail_def) lemma no_error : "good_context s \ snd (execute_instruction () s) = False" proof - assume "good_context s" hence "snd (execute_instruction() s) = False" using exe_instr_all by auto hence "snd (execute_instruction () s) = False" by (simp add: exec_ss2) thus ?thesis by assumption qed theorem single_step : "good_context s \ NEXT s = Some (snd (fst (execute_instruction () s)))" by (simp add: no_error next_match) (*********************************************************************) section \Privilege safty\ (*********************************************************************) text \The following shows that, if the pre-state is under user mode, then after a singel step execution, the post-state is aslo under user mode.\ lemma write_cpu_pc_privilege: "s' = snd (fst (write_cpu w PC s)) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: cpu_reg_mod_def) by (simp add: cpu_reg_val_def) lemma write_cpu_npc_privilege: "s' = snd (fst (write_cpu w nPC s)) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: cpu_reg_mod_def) by (simp add: cpu_reg_val_def) lemma write_cpu_y_privilege: "s' = snd (fst (write_cpu w Y s)) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: cpu_reg_mod_def) by (simp add: cpu_reg_val_def) lemma cpu_reg_mod_y_privilege: "s' = cpu_reg_mod w Y s \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" by (simp add: cpu_reg_mod_def cpu_reg_val_def) lemma cpu_reg_mod_asr_privilege: "s' = cpu_reg_mod w (ASR r) s \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" by (simp add: cpu_reg_mod_def cpu_reg_val_def) lemma global_reg_mod_privilege: "s' = global_reg_mod w1 n w2 s \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (induction n arbitrary:s) apply (clarsimp) apply (auto) apply (simp add: Let_def) by (simp add: cpu_reg_val_def) lemma out_reg_mod_privilege: "s' = out_reg_mod a w r s \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: out_reg_mod_def Let_def) by (simp add: cpu_reg_val_def) lemma in_reg_mod_privilege: "s' = in_reg_mod a w r s \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: in_reg_mod_def Let_def) by (simp add: cpu_reg_val_def) lemma user_reg_mod_privilege: -assumes a1: " s' = user_reg_mod d (w::(('a::len0) window_size)) r - (s::(('a::len0) sparc_state)) \ +assumes a1: " s' = user_reg_mod d (w::(('a::len) window_size)) r + (s::(('a::len) sparc_state)) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "r = 0") case True then show ?thesis using a1 by (simp add: user_reg_mod_def) next case False then have f1: "r \ 0" by auto then show ?thesis proof (cases "0 < r \ r < 8") case True then show ?thesis using a1 f1 apply (simp add: user_reg_mod_def) by (auto intro: global_reg_mod_privilege) next case False then have f2: "\(0 < r \ r < 8)" by auto then show ?thesis proof (cases "7 < r \ r < 16") case True then show ?thesis using a1 f1 f2 apply (simp add: user_reg_mod_def) by (auto intro: out_reg_mod_privilege) next case False then have f3: "\ (7 < r \ r < 16)" by auto then show ?thesis proof (cases "15 < r \ r < 24") case True then show ?thesis using a1 f1 f2 f3 apply (simp add: user_reg_mod_def) by (simp add: cpu_reg_val_def) next case False then show ?thesis using a1 f1 f2 f3 apply (simp add: user_reg_mod_def) by (auto intro: in_reg_mod_privilege) qed qed qed qed lemma write_reg_privilege: "s' = snd (fst (write_reg w1 w2 w3 - (s::(('a::len0) sparc_state)))) \ + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: write_reg_def simpler_modify_def) by (auto intro: user_reg_mod_privilege) lemma set_annul_privilege: "s' = snd (fst (set_annul b s)) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: set_annul_def simpler_modify_def) apply (simp add: annul_mod_def write_annul_def) by (simp add: cpu_reg_val_def) lemma set_reset_trap_privilege: "s' = snd (fst (set_reset_trap b s)) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: set_reset_trap_def simpler_modify_def) apply (simp add: reset_trap_mod_def write_annul_def) by (simp add: cpu_reg_val_def) lemma empty_delayed_pool_write_privilege: "get_delayed_pool s = [] \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ s' = delayed_pool_write s \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: delayed_pool_write_def) by (simp add: get_delayed_write_def delayed_write_all_def delayed_pool_rm_list_def) lemma raise_trap_privilege: "((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ s' = snd (fst (raise_trap t s)) \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: raise_trap_def) apply (simp add: simpler_modify_def add_trap_set_def) by (simp add: cpu_reg_val_def) lemma write_cpu_tt_privilege: "s' = snd (fst (write_cpu_tt w s)) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: write_cpu_tt_def) apply (simp add: exec_gets) apply (simp add: write_cpu_def cpu_reg_mod_def write_tt_def) apply (simp add: simpler_modify_def) by (simp add: cpu_reg_val_def) lemma emp_trap_set_privilege: "s' = emp_trap_set s \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: emp_trap_set_def) by (simp add: cpu_reg_val_def) lemma sys_reg_mod_privilege: "s' = sys_reg_mod w r s \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: sys_reg_mod_def) by (simp add: cpu_reg_val_def) lemma mem_mod_privilege: assumes a1: "s' = mem_mod a1 a2 v s \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "(uint a1) = 8 \ (uint a1) = 10") case True then show ?thesis using a1 apply (simp add: mem_mod_def) apply (simp add: Let_def) by (simp add: cpu_reg_val_def) next case False then have f1: "\((uint a1) = 8 \ (uint a1) = 10)" by auto then show ?thesis proof (cases "(uint a1) = 9 \ (uint a1) = 11") case True then show ?thesis using a1 f1 apply (simp add: mem_mod_def) apply (simp add: Let_def) by (simp add: cpu_reg_val_def) next case False then show ?thesis using a1 f1 apply (simp add: mem_mod_def) by (simp add: cpu_reg_val_def) qed qed lemma mem_mod_w32_privilege: "s' = mem_mod_w32 a1 a2 b d s \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: mem_mod_w32_def) apply (simp add: Let_def) by (auto intro: mem_mod_privilege) lemma add_instr_cache_privilege: "s' = add_instr_cache s addr y m \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: add_instr_cache_def) apply (simp add: Let_def) by (simp add: icache_mod_def cpu_reg_val_def) lemma add_data_cache_privilege: "s' = add_data_cache s addr y m \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: add_data_cache_def) apply (simp add: Let_def) by (simp add: dcache_mod_def cpu_reg_val_def) lemma memory_read_privilege: assumes a1: "s' = snd (memory_read asi addr s) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "uint asi = 1") case True then show ?thesis using a1 apply (simp add: memory_read_def) by (simp add: Let_def) next case False then have f1: "uint asi \ 1" by auto then show ?thesis proof (cases "uint asi = 2") case True then show ?thesis using a1 f1 by (simp add: memory_read_def) next case False then have f2: "uint asi \ 2" by auto then show ?thesis proof (cases "uint asi \ {8,9}") case True then have f3: "uint asi \ {8,9}" by auto then show ?thesis proof (cases "load_word_mem s addr asi = None") case True then have f4: "load_word_mem s addr asi = None" by auto then show ?thesis using a1 f1 f2 f3 f4 by (simp add: memory_read_def) next case False then show ?thesis using a1 f1 f2 f3 apply (simp add: memory_read_def) apply auto apply (simp add: add_instr_cache_privilege) by (simp add: add_instr_cache_privilege) qed next case False then have f5: "uint asi \ {8, 9}" by auto then show ?thesis proof (cases "uint asi \ {10,11}") case True then have f6: "uint asi \ {10,11}" by auto then show ?thesis proof (cases "load_word_mem s addr asi = None") case True then have f7: "load_word_mem s addr asi = None" by auto then show ?thesis using a1 f1 f2 f5 f6 f7 by (simp add: memory_read_def) next case False then show ?thesis using a1 f1 f2 f5 f6 apply (simp add: memory_read_def) apply auto apply (simp add: add_data_cache_privilege) by (simp add: add_data_cache_privilege) qed next case False then have f8: "uint asi \ {10,11}" by auto then show ?thesis proof (cases "uint asi = 13") case True then have f9: "uint asi = 13" by auto then show ?thesis proof (cases "read_instr_cache s addr = None") case True then show ?thesis using a1 f1 f2 f5 f8 f9 by (simp add: memory_read_def) next case False then show ?thesis using a1 f1 f2 f5 f8 f9 apply (simp add: memory_read_def) by auto qed next case False then have f10: "uint asi \ 13" by auto then show ?thesis proof (cases "uint asi = 15") case True then show ?thesis using a1 f1 f2 f5 f8 f10 apply (simp add: memory_read_def) apply (cases "read_data_cache s addr = None") by auto next case False then show ?thesis using a1 f1 f2 f5 f8 f10 apply (simp add: memory_read_def) \ \The rest cases are easy.\ by (simp add: Let_def) qed qed qed qed qed qed lemma get_curr_win_privilege: "s' = snd (fst (get_curr_win() s)) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: get_curr_win_def) by (simp add: simpler_gets_def) lemma load_sub2_privilege: assumes a1: "s' = snd (fst (load_sub2 addr asi r win w s)) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "fst (memory_read asi (addr + 4) (snd (fst (write_reg w win (r AND 30) s)))) = None") case True then show ?thesis using a1 apply (simp add: load_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (auto intro: raise_trap_privilege write_reg_privilege) next case False then show ?thesis using a1 apply (simp add: load_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply clarsimp apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) by (auto intro: write_reg_privilege memory_read_privilege) qed lemma load_sub3_privilege: assumes a1: "s' = snd (fst (load_sub3 instr curr_win rd asi address s)) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "fst (memory_read asi address s) = None") case True then show ?thesis using a1 apply (simp add: load_sub3_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) by (auto intro: raise_trap_privilege) next case False then have f1: "fst (memory_read asi address s) \ None " by auto then show ?thesis proof (cases "rd \ 0 \ (fst instr = load_store_type LD \ fst instr = load_store_type LDA \ fst instr = load_store_type LDUH \ fst instr = load_store_type LDSB \ fst instr = load_store_type LDUB \ fst instr = load_store_type LDUBA \ fst instr = load_store_type LDSH \ fst instr = load_store_type LDSHA \ fst instr = load_store_type LDUHA \ fst instr = load_store_type LDSBA)") case True then show ?thesis using a1 f1 apply (simp add: load_sub3_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply clarsimp apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) by (auto intro: write_reg_privilege memory_read_privilege) next case False then show ?thesis using a1 f1 apply (simp add: load_sub3_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply auto apply (simp add: simpler_modify_def bind_def h1_def h2_def) apply (auto intro: load_sub2_privilege memory_read_privilege) apply (simp add: simpler_modify_def bind_def h1_def h2_def) by (auto intro: load_sub2_privilege memory_read_privilege) qed qed lemma load_sub1_privilege: assumes a1: "s' = snd (fst (load_sub1 instr rd s_val s)) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: load_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply auto by (auto intro: get_curr_win_privilege raise_trap_privilege load_sub3_privilege) lemma load_instr_privilege: "s' = snd (fst (load_instr i s)) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: load_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) apply clarsimp by (auto intro: get_curr_win_privilege raise_trap_privilege load_sub1_privilege) lemma store_barrier_pending_mod_privilege: "s' = store_barrier_pending_mod b s \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: store_barrier_pending_mod_def) apply (simp add: write_store_barrier_pending_def) by (simp add: cpu_reg_val_def) lemma store_word_mem_privilege: assumes a1: "store_word_mem s addr data byte_mask asi = Some s' \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: store_word_mem_def) apply (case_tac "virt_to_phys addr (mmu s) (mem s) = None") apply auto apply (case_tac "mmu_writable (get_acc_flag b) asi") apply auto by (simp add: mem_mod_w32_privilege) lemma flush_instr_cache_privilege: "((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ s' = flush_instr_cache s \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: flush_instr_cache_def) by (simp add: cpu_reg_val_def) lemma flush_data_cache_privilege: "((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ s' = flush_data_cache s \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: flush_data_cache_def) by (simp add: cpu_reg_val_def) lemma flush_cache_all_privilege: "((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0 \ s' = flush_cache_all s \ ((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: flush_cache_all_def) by (simp add: cpu_reg_val_def) lemma memory_write_asi_privilege: assumes a1: "r = memory_write_asi asi addr byte_mask data s \ r = Some s' \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "uint asi = 1") case True then show ?thesis using a1 apply (simp add: memory_write_asi_def) by (auto intro: store_word_mem_privilege) next case False then have f1: "uint asi \ 1" by auto then show ?thesis proof (cases "uint asi = 2") case True then have f01: "uint asi = 2" by auto then show ?thesis proof (cases "uint addr = 0") case True then show ?thesis using a1 f1 f01 apply (simp add: memory_write_asi_def) apply (simp add: ccr_flush_def) apply (simp add: Let_def) apply auto apply (metis flush_data_cache_privilege flush_instr_cache_privilege sys_reg_mod_privilege) apply (metis flush_instr_cache_privilege sys_reg_mod_privilege) apply (metis flush_data_cache_privilege sys_reg_mod_privilege) by (simp add: sys_reg_mod_privilege) next case False then show ?thesis using a1 f1 f01 apply (simp add: memory_write_asi_def) apply clarsimp by (metis option.distinct(1) option.sel sys_reg_mod_privilege) qed next case False then have f2: "uint asi \ 2" by auto then show ?thesis proof (cases "uint asi \ {8,9}") case True then show ?thesis using a1 f1 f2 apply (simp add: memory_write_asi_def) using store_word_mem_privilege add_instr_cache_privilege by blast next case False then have f3: "uint asi \ {8,9}" by auto then show ?thesis proof (cases "uint asi \ {10,11}") case True then show ?thesis using a1 f1 f2 f3 apply (simp add: memory_write_asi_def) using store_word_mem_privilege add_data_cache_privilege by blast next case False then have f4: "uint asi \ {10,11}" by auto then show ?thesis proof (cases "uint asi = 13") case True then show ?thesis using a1 f1 f2 f3 f4 apply (simp add: memory_write_asi_def) by (auto simp add: add_instr_cache_privilege) next case False then have f5: "uint asi \ 13" by auto then show ?thesis proof (cases "uint asi = 15") case True then show ?thesis using a1 f1 f2 f3 f4 f5 apply (simp add: memory_write_asi_def) by (auto simp add: add_data_cache_privilege) next case False then have f6: "uint asi \ 15" by auto then show ?thesis proof (cases "uint asi = 16") case True then show ?thesis using a1 apply (simp add: memory_write_asi_def) by (auto simp add: flush_instr_cache_privilege) next case False then have f7: "uint asi \ 16" by auto then show ?thesis proof (cases "uint asi = 17") case True then show ?thesis using a1 apply (simp add: memory_write_asi_def) by (auto simp add: flush_data_cache_privilege) next case False then have f8: "uint asi \ 17" by auto then show ?thesis proof (cases "uint asi = 24") case True then show ?thesis using a1 apply (simp add: memory_write_asi_def) by (auto simp add: flush_cache_all_privilege) next case False then have f9: "uint asi \ 24" by auto then show ?thesis proof (cases "uint asi = 25") case True then show ?thesis using a1 apply (simp add: memory_write_asi_def) apply (case_tac "mmu_reg_mod (mmu s) addr data = None") apply auto by (simp add: cpu_reg_val_def) next case False then have f10: "uint asi \ 25" by auto then show ?thesis proof (cases "uint asi = 28") case True then show ?thesis using a1 apply (simp add: memory_write_asi_def) by (auto simp add: mem_mod_w32_privilege) next case False \ \The remaining cases are easy.\ then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 apply (simp add: memory_write_asi_def) apply (auto simp add: Let_def) apply (case_tac "uint asi = 20 \ uint asi = 21") by auto qed qed qed qed qed qed qed qed qed qed qed lemma memory_write_privilege: assumes a1: "r = memory_write asi addr byte_mask data - (s::(('a::len0) sparc_state)) \ + (s::(('a::len) sparc_state)) \ r = Some s' \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR - (s'::(('a::len0) sparc_state)))))::word1) = 0" + (s'::(('a::len) sparc_state)))))::word1) = 0" proof - have "\x. Some x \ None" by auto then have "r \ None" using a1 by (simp add: \r = memory_write asi addr byte_mask data s \ r = Some s' \ ucast (get_S (cpu_reg_val PSR s)) = 0\) then have "\s''. r = Some (store_barrier_pending_mod False s'')" using a1 by (metis (no_types, lifting) memory_write_def option.case_eq_if) then have "\s''. s' = store_barrier_pending_mod False s''" using a1 by blast then have "\s''. memory_write_asi asi addr byte_mask data s = Some s'' \ s' = store_barrier_pending_mod False s''" by (metis (no_types, lifting) assms memory_write_def not_None_eq option.case_eq_if option.sel) then show ?thesis using a1 using memory_write_asi_privilege store_barrier_pending_mod_privilege by blast qed lemma store_sub2_privilege: assumes a1: "s' = snd (fst (store_sub2 instr curr_win rd asi address s)) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "memory_write asi address (st_byte_mask instr address) (st_data0 instr curr_win rd address s) s = None") case True then show ?thesis using a1 apply (simp add: store_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (metis fst_conv raise_trap_privilege return_def snd_conv) next case False then have f1: "\(memory_write asi address (st_byte_mask instr address) (st_data0 instr curr_win rd address s) s = None)" by auto then show ?thesis proof (cases "(fst instr) \ {load_store_type STD,load_store_type STDA}") case True then have f2: "(fst instr) \ {load_store_type STD,load_store_type STDA}" by auto then show ?thesis using a1 f1 apply (simp add: store_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) apply (simp add: return_def) apply (simp add: bind_def case_prod_unfold) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: case_prod_unfold bind_def h1_def h2_def Let_def simpler_modify_def) apply (simp add: simpler_gets_def) apply auto using memory_write_privilege raise_trap_privilege apply blast apply (simp add: simpler_modify_def simpler_gets_def bind_def) apply (meson memory_write_privilege) using memory_write_privilege raise_trap_privilege apply blast by (meson memory_write_privilege) next case False then show ?thesis using a1 f1 apply (simp add: store_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply clarsimp apply (simp add: simpler_modify_def return_def) by (auto intro: memory_write_privilege) qed qed lemma store_sub1_privilege: assumes a1: "s' = snd (fst (store_sub1 instr rd s_val - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR - (s'::(('a::len0) sparc_state)))))::word1) = 0" + (s'::(('a::len) sparc_state)))))::word1) = 0" proof (cases "(fst instr = load_store_type STH \ fst instr = load_store_type STHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s)))))::word1) \ 0") case True then show ?thesis using a1 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) using get_curr_win_privilege raise_trap_privilege by blast next case False then have f1: "\((fst instr = load_store_type STH \ fst instr = load_store_type STHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s)))))::word1) \ 0)" by auto then show ?thesis proof (cases "(fst instr \ {load_store_type ST,load_store_type STA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s)))))::word2) \ 0") case True then show ?thesis using a1 f1 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) using get_curr_win_privilege raise_trap_privilege by blast next case False then have f2: "\((fst instr \ {load_store_type ST,load_store_type STA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s)))))::word2) \ 0)" by auto then show ?thesis proof (cases "(fst instr \ {load_store_type STD,load_store_type STDA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s)))))::word3) \ 0") case True then show ?thesis using a1 f1 f2 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) using get_curr_win_privilege raise_trap_privilege by blast next case False then show ?thesis using a1 f1 f2 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (meson get_curr_win_privilege store_sub2_privilege) qed qed qed lemma store_instr_privilege: assumes a1: "s' = snd (fst (store_instr instr - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR - (s'::(('a::len0) sparc_state)))))::word1) = 0" + (s'::(('a::len) sparc_state)))))::word1) = 0" using a1 apply (simp add: store_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) using raise_trap_privilege store_sub1_privilege by blast lemma sethi_instr_privilege: assumes a1: "s' = snd (fst (sethi_instr instr - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: sethi_instr_def) apply (simp add: Let_def) apply auto apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) using get_curr_win_privilege write_reg_privilege apply blast by (simp add: return_def) lemma nop_instr_privilege: assumes a1: "s' = snd (fst (nop_instr instr - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: nop_instr_def) by (simp add: return_def) lemma ucast_0: "((ucast (get_S w))::word1) = 0 \ get_S w = 0" by (simp add: ucast_id) lemma ucast_02: "get_S w = 0 \ ((ucast (get_S w))::word1) = 0" by simp lemma ucast_s: "((ucast (get_S w))::word1) = 0 \ (AND) w (0b00000000000000000000000010000000::word32) = 0" apply (simp add: get_S_def) by (metis (mono_tags) ucast_id zero_neq_one) lemma ucast_s2: "(AND) w 0b00000000000000000000000010000000 = 0 \ ((ucast (get_S w))::word1) = 0" by (simp add: get_S_def) lemma update_PSR_icc_1: "w' = (AND) w (0b11111111000011111111111111111111::word32) \ ((ucast (get_S w))::word1) = 0 \ ((ucast (get_S w'))::word1) = 0" by (simp add: get_S_def word_bw_assocs(1)) lemma and_num_1048576_128: "(AND) (0b00000000000100000000000000000000::word32) (0b00000000000000000000000010000000::word32) = 0" by simp lemma and_num_2097152_128: "(AND) (0b00000000001000000000000000000000::word32) (0b00000000000000000000000010000000::word32) = 0" by simp lemma and_num_4194304_128: "(AND) (0b00000000010000000000000000000000::word32) (0b00000000000000000000000010000000::word32) = 0" by simp lemma and_num_8388608_128: "(AND) (0b00000000100000000000000000000000::word32) (0b00000000000000000000000010000000::word32) = 0" by simp lemma or_and_s: "(AND) w1 (0b00000000000000000000000010000000::word32) = 0 \ (AND) w2 (0b00000000000000000000000010000000::word32) = 0 \ (AND) ((OR) w1 w2) (0b00000000000000000000000010000000::word32) = 0" by (simp add: word_ao_dist) lemma and_or_s: assumes a1: "((ucast (get_S w1))::word1) = 0 \ (AND) w2 (0b00000000000000000000000010000000::word32) = 0" shows "((ucast (get_S ((OR) ((AND) w1 (0b11111111000011111111111111111111::word32)) w2)))::word1) = 0" by (metis (full_types) assms ucast_s ucast_s2 word_ao_absorbs(8) word_bool_alg.conj_disj_distrib2) lemma and_or_or_s: assumes a1: "((ucast (get_S w1))::word1) = 0 \ (AND) w2 (0b00000000000000000000000010000000::word32) = 0 \ (AND) w3 (0b00000000000000000000000010000000::word32) = 0" shows "((ucast (get_S ((OR) ((OR) ((AND) w1 (0b11111111000011111111111111111111::word32)) w2) w3)))::word1) = 0" using and_or_s assms or_and_s ucast_s ucast_s2 by blast lemma and_or_or_or_s: assumes a1: "((ucast (get_S w1))::word1) = 0 \ (AND) w2 (0b00000000000000000000000010000000::word32) = 0 \ (AND) w3 (0b00000000000000000000000010000000::word32) = 0 \ (AND) w4 (0b00000000000000000000000010000000::word32) = 0" shows "((ucast (get_S ((OR) ((OR) ((OR) ((AND) w1 (0b11111111000011111111111111111111::word32)) w2) w3) w4)))::word1) = 0" using and_or_or_s assms or_and_s ucast_s ucast_s2 by (meson word_bool_alg.conj.commute word_bool_alg.conj_zero_left word_bw_assocs(1)) lemma and_or_or_or_or_s: assumes a1: "((ucast (get_S w1))::word1) = 0 \ (AND) w2 (0b00000000000000000000000010000000::word32) = 0 \ (AND) w3 (0b00000000000000000000000010000000::word32) = 0 \ (AND) w4 (0b00000000000000000000000010000000::word32) = 0 \ (AND) w5 (0b00000000000000000000000010000000::word32) = 0" shows "((ucast (get_S ((OR) ((OR) ((OR) ((OR) ((AND) w1 (0b11111111000011111111111111111111::word32)) w2) w3) w4) w5)))::word1) = 0" using and_or_or_or_s assms or_and_s ucast_s ucast_s2 by (meson word_ao_absorbs(8) word_bool_alg.conj_disj_distrib2) lemma write_cpu_PSR_icc_privilege: assumes a1: "s' = snd (fst (write_cpu (update_PSR_icc n_val z_val v_val c_val (cpu_reg_val PSR s)) PSR - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: write_cpu_def) apply (simp add: simpler_modify_def) apply (simp add: cpu_reg_mod_def update_PSR_icc_def) apply (simp add: cpu_reg_val_def) apply auto using update_PSR_icc_1 apply blast using update_PSR_icc_1 and_num_1048576_128 and_or_s apply blast using update_PSR_icc_1 and_num_2097152_128 and_or_s apply blast using update_PSR_icc_1 and_num_1048576_128 and_num_2097152_128 and_or_or_s apply blast using update_PSR_icc_1 and_num_4194304_128 and_or_s apply blast using update_PSR_icc_1 and_num_1048576_128 and_num_4194304_128 and_or_or_s apply blast using update_PSR_icc_1 and_num_2097152_128 and_num_4194304_128 and_or_or_s apply blast using update_PSR_icc_1 and_num_1048576_128 and_num_2097152_128 and_num_4194304_128 and_or_or_or_s apply blast using update_PSR_icc_1 and_num_8388608_128 and_or_s apply blast using update_PSR_icc_1 and_num_1048576_128 and_num_8388608_128 and_or_or_s apply blast using update_PSR_icc_1 and_num_2097152_128 and_num_8388608_128 and_or_or_s apply blast using update_PSR_icc_1 and_num_1048576_128 and_num_2097152_128 and_num_8388608_128 and_or_or_or_s apply blast using update_PSR_icc_1 and_num_4194304_128 and_num_8388608_128 and_or_or_s apply blast using update_PSR_icc_1 and_num_1048576_128 and_num_4194304_128 and_num_8388608_128 and_or_or_or_s apply blast using update_PSR_icc_1 and_num_2097152_128 and_num_4194304_128 and_num_8388608_128 and_or_or_or_s apply blast using update_PSR_icc_1 and_num_1048576_128 and_num_2097152_128 and_num_4194304_128 and_num_8388608_128 and_or_or_or_or_s by blast lemma and_num_4294967167_128: "(AND) (0b11111111111111111111111101111111::word32) (0b00000000000000000000000010000000::word32) = 0" by simp lemma s_0_word: "((ucast (get_S ((AND) w (0b11111111111111111111111101111111::word32))))::word1) = 0" apply (simp add: get_S_def) using and_num_4294967167_128 by (simp add: word_bool_alg.conj.commute word_bw_lcs(1)) lemma update_PSR_CWP_1: "w' = (AND) w (0b11111111111111111111111111100000::word32) \ ((ucast (get_S w))::word1) = 0 \ ((ucast (get_S w'))::word1) = 0" by (simp add: get_S_def word_bw_assocs(1)) lemma write_cpu_PSR_CWP_privilege: assumes a1: "s' = snd (fst (write_cpu (update_CWP cwp_val (cpu_reg_val PSR s)) PSR - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: write_cpu_def) apply (simp add: simpler_modify_def) apply (simp add: cpu_reg_mod_def) apply (simp add: update_CWP_def) apply (simp add: Let_def) apply auto apply (simp add: cpu_reg_val_def) using s_0_word by blast lemma logical_instr_sub1_privilege: assumes a1: "s' = snd (fst (logical_instr_sub1 instr_name result - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "instr_name = logic_type ANDcc \ instr_name = logic_type ANDNcc \ instr_name = logic_type ORcc \ instr_name = logic_type ORNcc \ instr_name = logic_type XORcc \ instr_name = logic_type XNORcc") case True then show ?thesis using a1 apply (simp add: logical_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: logical_new_psr_val_def) using write_cpu_PSR_icc_privilege by blast next case False then show ?thesis using a1 apply (simp add: logical_instr_sub1_def) by (simp add: return_def) qed lemma logical_instr_privilege: assumes a1: "s' = snd (fst (logical_instr instr - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: logical_instr_def) apply (simp add: Let_def simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply auto apply (meson get_curr_win_privilege logical_instr_sub1_privilege write_reg_privilege) by (meson get_curr_win_privilege logical_instr_sub1_privilege write_reg_privilege) method shift_instr_privilege_proof = ( (simp add: shift_instr_def), (simp add: Let_def), (simp add: simpler_gets_def), (simp add: bind_def h1_def h2_def Let_def case_prod_unfold), auto, (blast intro: get_curr_win_privilege write_reg_privilege), (blast intro: get_curr_win_privilege write_reg_privilege) ) lemma shift_instr_privilege: assumes a1: "s' = snd (fst (shift_instr instr - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "(fst instr = shift_type SLL) \ (get_operand_w5 ((snd instr)!3) \ 0)") case True then show ?thesis using a1 by shift_instr_privilege_proof next case False then have f1: "\((fst instr = shift_type SLL) \ (get_operand_w5 ((snd instr)!3) \ 0))" by auto then show ?thesis proof (cases "(fst instr = shift_type SRL) \ (get_operand_w5 ((snd instr)!3) \ 0)") case True then show ?thesis using a1 f1 by shift_instr_privilege_proof next case False then have f2: "\((fst instr = shift_type SRL) \ (get_operand_w5 ((snd instr)!3) \ 0))" by auto then show ?thesis proof (cases "(fst instr = shift_type SRA) \ (get_operand_w5 ((snd instr)!3) \ 0)") case True then show ?thesis using a1 f1 f2 by shift_instr_privilege_proof next case False then show ?thesis using a1 f1 f2 apply (simp add: shift_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def) apply (simp add: bind_def h1_def h2_def Let_def case_prod_unfold) apply (simp add: return_def) using get_curr_win_privilege by blast qed qed qed lemma add_instr_sub1_privilege: assumes a1: "s' = snd (fst (add_instr_sub1 instr_name result rs1_val operand2 - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "instr_name = arith_type ADDcc \ instr_name = arith_type ADDXcc") case True then show ?thesis using a1 apply (simp add: add_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (blast intro: write_cpu_PSR_icc_privilege) next case False then show ?thesis using a1 apply (simp add: add_instr_sub1_def) by (simp add: return_def) qed lemma add_instr_privilege: assumes a1: "s' = snd (fst (add_instr instr - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: add_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (meson add_instr_sub1_privilege get_curr_win_privilege write_reg_privilege) lemma sub_instr_sub1_privilege: assumes a1: "s' = snd (fst (sub_instr_sub1 instr_name result rs1_val operand2 - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "instr_name = arith_type SUBcc \ instr_name = arith_type SUBXcc") case True then show ?thesis using a1 apply (simp add: sub_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (blast intro: write_cpu_PSR_icc_privilege) next case False then show ?thesis using a1 apply (simp add: sub_instr_sub1_def) by (simp add: return_def) qed lemma sub_instr_privilege: assumes a1: "s' = snd (fst (sub_instr instr - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: sub_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (meson sub_instr_sub1_privilege get_curr_win_privilege write_reg_privilege) lemma mul_instr_sub1_privilege: assumes a1: "s' = snd (fst (mul_instr_sub1 instr_name result - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "instr_name \ {arith_type SMULcc,arith_type UMULcc}") case True then show ?thesis using a1 apply (simp add: mul_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (blast intro: write_cpu_PSR_icc_privilege) next case False then show ?thesis using a1 apply (simp add: mul_instr_sub1_def) by (simp add: return_def) qed lemma mul_instr_privilege: assumes a1: "s' = snd (fst (mul_instr instr - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: mul_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (meson get_curr_win_privilege mul_instr_sub1_privilege write_cpu_y_privilege write_reg_privilege) lemma div_write_new_val_privilege: assumes a1: "s' = snd (fst (div_write_new_val i result temp_V - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "(fst i) \ {arith_type UDIVcc,arith_type SDIVcc}") case True then show ?thesis using a1 apply (simp add: div_write_new_val_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (blast intro: write_cpu_PSR_icc_privilege) next case False then show ?thesis using a1 apply (simp add: div_write_new_val_def) by (simp add: return_def) qed lemma div_comp_privilege: assumes a1: "s' = snd (fst (div_comp instr rs1 rd operand2 - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: div_comp_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (meson get_curr_win_privilege div_write_new_val_privilege write_reg_privilege) lemma div_instr_privilege: assumes a1: "s' = snd (fst (div_instr instr - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: div_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: return_def) apply auto using raise_trap_privilege apply blast using div_comp_privilege by blast lemma save_retore_sub1_privilege: assumes a1: "s' = snd (fst (save_retore_sub1 result new_cwp rd - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: save_retore_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) using write_cpu_PSR_CWP_privilege write_reg_privilege by blast method save_restore_instr_privilege_proof = ( (simp add: save_restore_instr_def), (simp add: Let_def), (simp add: simpler_gets_def bind_def h1_def h2_def Let_def), (simp add: case_prod_unfold), auto, (blast intro: get_curr_win_privilege raise_trap_privilege), (simp add: simpler_gets_def bind_def h1_def h2_def Let_def case_prod_unfold), (blast intro: get_curr_win_privilege save_retore_sub1_privilege) ) lemma save_restore_instr_privilege: assumes a1: "s' = snd (fst (save_restore_instr instr - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "fst instr = ctrl_type SAVE") case True then have f1: "fst instr = ctrl_type SAVE" by auto then show ?thesis using a1 by save_restore_instr_privilege_proof next case False then show ?thesis using a1 by save_restore_instr_privilege_proof qed lemma call_instr_privilege: assumes a1: "s' = snd (fst (call_instr instr - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: call_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (meson get_curr_win_privilege write_cpu_npc_privilege write_cpu_pc_privilege write_reg_privilege) lemma jmpl_instr_privilege: assumes a1: "s' = snd (fst (jmpl_instr instr - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: jmpl_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply auto using get_curr_win_privilege raise_trap_privilege apply blast apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (meson get_curr_win_privilege write_cpu_npc_privilege write_cpu_pc_privilege write_reg_privilege) lemma rett_instr_privilege: assumes a1: "snd (rett_instr i s) = False \ s' = snd (fst (rett_instr instr - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: rett_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply auto apply (simp add: case_prod_unfold) apply (simp add: return_def) apply (blast intro: raise_trap_privilege) apply (simp add: bind_def h1_def h2_def Let_def) by (simp add: case_prod_unfold fail_def) method read_state_reg_instr_privilege_proof = ( (simp add: read_state_reg_instr_def), (simp add: Let_def), (simp add: simpler_gets_def bind_def h1_def h2_def Let_def), (simp add: case_prod_unfold) ) lemma read_state_reg_instr_privilege: assumes a1: "s' = snd (fst (read_state_reg_instr instr - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "(fst instr \ {sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR} \ (fst instr = sreg_type RDASR \ privileged_ASR (get_operand_w5 ((snd instr)!0))))") case True then have "(fst instr \ {sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR} \ (fst instr = sreg_type RDASR \ privileged_ASR (get_operand_w5 ((snd instr)!0)))) \ ((ucast (get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s))))))::word1) = 0" by (metis assms get_curr_win_privilege) then show ?thesis using a1 apply read_state_reg_instr_privilege_proof by (blast intro: raise_trap_privilege get_curr_win_privilege) next case False then have f1: "\((fst instr = sreg_type RDPSR \ fst instr = sreg_type RDWIM \ fst instr = sreg_type RDTBR \ fst instr = sreg_type RDASR \ privileged_ASR (get_operand_w5 (snd instr ! 0))) \ ucast (get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s))))) = 0)" by blast then show ?thesis proof (cases "illegal_instruction_ASR (get_operand_w5 ((snd instr)!0))") case True then show ?thesis using a1 f1 apply read_state_reg_instr_privilege_proof by (simp add: illegal_instruction_ASR_def) next case False then have f2: "\(illegal_instruction_ASR (get_operand_w5 ((snd instr)!0)))" by auto then show ?thesis proof (cases "(get_operand_w5 ((snd instr)!1)) \ 0") case True then have f3: "(get_operand_w5 ((snd instr)!1)) \ 0" by auto then show ?thesis proof (cases "fst instr = sreg_type RDY") case True then show ?thesis using a1 f1 f2 f3 apply (simp add: read_state_reg_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (blast intro: get_curr_win_privilege write_reg_privilege) next case False then have f4: "\(fst instr = sreg_type RDY)" by auto then show ?thesis proof (cases "fst instr = sreg_type RDASR") case True then show ?thesis using a1 f1 f2 f3 f4 apply read_state_reg_instr_privilege_proof apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (blast intro: get_curr_win_privilege write_reg_privilege) next case False then have f5: "\(fst instr = sreg_type RDASR)" by auto then show ?thesis proof (cases "fst instr = sreg_type RDPSR") case True then show ?thesis using a1 f1 f2 f3 f4 f5 apply read_state_reg_instr_privilege_proof by (blast intro: get_curr_win_privilege write_reg_privilege) next case False then show ?thesis using a1 f1 f2 f3 f4 f5 apply read_state_reg_instr_privilege_proof apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (blast intro: get_curr_win_privilege write_reg_privilege) qed qed qed next case False then show ?thesis using a1 apply read_state_reg_instr_privilege_proof apply (simp add: return_def) using f1 f2 get_curr_win_privilege by blast qed qed qed method write_state_reg_instr_privilege_proof = ( (simp add: write_state_reg_instr_def), (simp add: Let_def), (simp add: simpler_gets_def bind_def h1_def h2_def Let_def), (simp add: case_prod_unfold) ) lemma write_state_reg_instr_privilege: assumes a1: "s' = snd (fst (write_state_reg_instr instr - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "fst instr = sreg_type WRY") case True then show ?thesis using a1 apply write_state_reg_instr_privilege_proof apply (simp add: simpler_modify_def) apply (simp add: delayed_pool_add_def DELAYNUM_def) by (blast intro: cpu_reg_mod_y_privilege get_curr_win_privilege) next case False then have f1: "\(fst instr = sreg_type WRY)" by auto then show ?thesis proof (cases "fst instr = sreg_type WRASR") case True then show ?thesis using a1 f1 apply write_state_reg_instr_privilege_proof apply (simp add: simpler_modify_def) apply auto using illegal_instruction_ASR_def apply blast using illegal_instruction_ASR_def apply blast using illegal_instruction_ASR_def apply blast using raise_trap_privilege get_curr_win_privilege apply blast apply (simp add: simpler_modify_def delayed_pool_add_def DELAYNUM_def) using cpu_reg_mod_asr_privilege get_curr_win_privilege apply blast apply (simp add: simpler_modify_def delayed_pool_add_def DELAYNUM_def) using cpu_reg_mod_asr_privilege get_curr_win_privilege by blast next case False then have f2: "\(fst instr = sreg_type WRASR)" by auto have f3: "get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s)))) = 0" using get_curr_win_privilege a1 by (metis ucast_id) then show ?thesis proof (cases "fst instr = sreg_type WRPSR") case True then show ?thesis using a1 f1 f2 f3 apply write_state_reg_instr_privilege_proof by (metis raise_trap_privilege ucast_0) next case False then have f4: "\(fst instr = sreg_type WRPSR)" by auto then show ?thesis proof (cases "fst instr = sreg_type WRWIM") case True then show ?thesis using a1 f1 f2 f3 f4 apply write_state_reg_instr_privilege_proof by (metis raise_trap_privilege ucast_0) next case False then have f5: "\(fst instr = sreg_type WRWIM)" by auto then show ?thesis using a1 f1 f2 f3 f4 f5 apply write_state_reg_instr_privilege_proof by (metis raise_trap_privilege ucast_0) qed qed qed qed lemma flush_instr_privilege: assumes a1: "s' = snd (fst (flush_instr instr - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: flush_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def simpler_modify_def) by (auto simp add: flush_cache_all_privilege) lemma branch_instr_privilege: assumes a1: "s' = snd (fst (branch_instr instr - (s::(('a::len0) sparc_state)))) + (s::(('a::len) sparc_state)))) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: branch_instr_def) apply (simp add: Let_def simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold return_def) by (meson set_annul_privilege write_cpu_npc_privilege write_cpu_pc_privilege) method dispath_instr_privilege_proof = ( (simp add: dispatch_instruction_def), (simp add: simpler_gets_def bind_def h1_def h2_def Let_def), (simp add: Let_def) ) lemma dispath_instr_privilege: assumes a1: "snd (dispatch_instruction instr s) = False \ s' = snd (fst (dispatch_instruction instr s)) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "get_trap_set s = {}") case True then have f1: "get_trap_set s = {}" by auto show ?thesis proof (cases "fst instr \ {load_store_type LDSB,load_store_type LDUB, load_store_type LDUBA,load_store_type LDUH,load_store_type LD, load_store_type LDA,load_store_type LDD}") case True then show ?thesis using a1 f1 apply dispath_instr_privilege_proof by (blast intro: load_instr_privilege) next case False then have f2: "\(fst instr \ {load_store_type LDSB,load_store_type LDUB, load_store_type LDUBA,load_store_type LDUH,load_store_type LD, load_store_type LDA,load_store_type LDD})" by auto then show ?thesis proof (cases "fst instr \ {load_store_type STB,load_store_type STH, load_store_type ST,load_store_type STA,load_store_type STD}") case True then show ?thesis using a1 f1 f2 apply dispath_instr_privilege_proof by (blast intro: store_instr_privilege) next case False then have f3: "\(fst instr \ {load_store_type STB,load_store_type STH, load_store_type ST,load_store_type STA,load_store_type STD})" by auto then show ?thesis proof (cases "fst instr \ {sethi_type SETHI}") case True then show ?thesis using a1 f1 f2 f3 apply dispath_instr_privilege_proof by (blast intro: sethi_instr_privilege) next case False then have f4: "\(fst instr \ {sethi_type SETHI})" by auto then show ?thesis proof (cases "fst instr \ {nop_type NOP}") case True then show ?thesis using a1 f1 f2 f3 f4 apply dispath_instr_privilege_proof by (blast intro: nop_instr_privilege) next case False then have f5: "\(fst instr \ {nop_type NOP})" by auto then show ?thesis proof (cases "fst instr \ {logic_type ANDs,logic_type ANDcc,logic_type ANDN, logic_type ANDNcc,logic_type ORs,logic_type ORcc,logic_type ORN, logic_type XORs,logic_type XNOR}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 apply dispath_instr_privilege_proof by (blast intro: logical_instr_privilege) next case False then have f6: "\(fst instr \ {logic_type ANDs,logic_type ANDcc,logic_type ANDN, logic_type ANDNcc,logic_type ORs,logic_type ORcc,logic_type ORN, logic_type XORs,logic_type XNOR})" by auto show ?thesis proof (cases "fst instr \ {shift_type SLL,shift_type SRL,shift_type SRA}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 apply dispath_instr_privilege_proof by (blast intro: shift_instr_privilege) next case False then have f7: "\(fst instr \ {shift_type SLL,shift_type SRL,shift_type SRA})" by auto then show ?thesis proof (cases "fst instr \ {arith_type ADD,arith_type ADDcc,arith_type ADDX}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 apply dispath_instr_privilege_proof by (blast intro: add_instr_privilege) next case False then have f8: "\(fst instr \ {arith_type ADD,arith_type ADDcc,arith_type ADDX})" by auto then show ?thesis proof (cases "fst instr \ {arith_type SUB,arith_type SUBcc,arith_type SUBX}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 apply dispath_instr_privilege_proof by (blast intro: sub_instr_privilege) next case False then have f9: "\(fst instr \ {arith_type SUB,arith_type SUBcc,arith_type SUBX})" by auto then show ?thesis proof (cases "fst instr \ {arith_type UMUL,arith_type SMUL,arith_type SMULcc}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 apply dispath_instr_privilege_proof by (blast intro: mul_instr_privilege) next case False then have f10: "\(fst instr \ {arith_type UMUL,arith_type SMUL, arith_type SMULcc})" by auto then show ?thesis proof (cases "fst instr \ {arith_type UDIV,arith_type UDIVcc,arith_type SDIV}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 apply dispath_instr_privilege_proof by (blast intro: div_instr_privilege) next case False then have f11: "\(fst instr \ {arith_type UDIV, arith_type UDIVcc,arith_type SDIV})" by auto then show ?thesis proof (cases "fst instr \ {ctrl_type SAVE,ctrl_type RESTORE}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 apply dispath_instr_privilege_proof by (blast intro: save_restore_instr_privilege) next case False then have f12: "\(fst instr \ {ctrl_type SAVE,ctrl_type RESTORE})" by auto then show ?thesis proof (cases "fst instr \ {call_type CALL}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 apply dispath_instr_privilege_proof by (blast intro: call_instr_privilege) next case False then have f13: "\(fst instr \ {call_type CALL})" by auto then show ?thesis proof (cases "fst instr \ {ctrl_type JMPL}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 apply dispath_instr_privilege_proof by (blast intro: jmpl_instr_privilege) next case False then have f14: "\(fst instr \ {ctrl_type JMPL})" by auto then show ?thesis proof (cases "fst instr \ {ctrl_type RETT}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 apply dispath_instr_privilege_proof by (blast intro: rett_instr_privilege) next case False then have f15: "\(fst instr \ {ctrl_type RETT})" by auto then show ?thesis proof (cases "fst instr \ {sreg_type RDY,sreg_type RDPSR, sreg_type RDWIM, sreg_type RDTBR}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 apply dispath_instr_privilege_proof by (blast intro: read_state_reg_instr_privilege) next case False then have f16: "\(fst instr \ {sreg_type RDY,sreg_type RDPSR, sreg_type RDWIM, sreg_type RDTBR})" by auto then show ?thesis proof (cases "fst instr \ {sreg_type WRY,sreg_type WRPSR, sreg_type WRWIM, sreg_type WRTBR}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 apply dispath_instr_privilege_proof by (blast intro: write_state_reg_instr_privilege) next case False then have f17: "\(fst instr \ {sreg_type WRY,sreg_type WRPSR, sreg_type WRWIM, sreg_type WRTBR})" by auto then show ?thesis proof (cases "fst instr \ {load_store_type FLUSH}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 apply dispath_instr_privilege_proof by (blast intro: flush_instr_privilege) next case False then have f18: "\(fst instr \ {load_store_type FLUSH})" by auto then show ?thesis proof (cases "fst instr \ {bicc_type BE,bicc_type BNE, bicc_type BGU,bicc_type BLE,bicc_type BL,bicc_type BGE, bicc_type BNEG,bicc_type BG,bicc_type BCS,bicc_type BLEU, bicc_type BCC,bicc_type BA,bicc_type BN}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 apply dispath_instr_privilege_proof by (blast intro: branch_instr_privilege) next case False then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 apply dispath_instr_privilege_proof by (simp add: fail_def) qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed next case False then show ?thesis using a1 apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) by (simp add: return_def) qed lemma execute_instr_sub1_privilege: assumes a1: "snd (execute_instr_sub1 i s) = False \ s' = snd (fst (execute_instr_sub1 i s)) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "get_trap_set s = {} \ fst i \ {call_type CALL,ctrl_type RETT,ctrl_type JMPL, bicc_type BE,bicc_type BNE,bicc_type BGU, bicc_type BLE,bicc_type BL,bicc_type BGE, bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC, bicc_type BA,bicc_type BN}") case True then show ?thesis using a1 apply (simp add: execute_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold return_def) by (auto intro: write_cpu_pc_privilege write_cpu_npc_privilege) next case False then show ?thesis using a1 apply (simp add: execute_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold return_def) by auto qed text \ Assume that there is no \delayed_write\ and there is no traps to be executed. If an instruction is executed as a user, the privilege will not be changed to supervisor after the execution.\ theorem safe_privilege : assumes a1: "get_delayed_pool s = [] \ get_trap_set s = {} \ snd (execute_instruction() s) = False \ s' = snd (fst (execute_instruction() s)) \ ((ucast (get_S (cpu_reg_val PSR s)))::word1) = 0" shows "((ucast (get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "exe_mode_val s") case True then have f2: "exe_mode_val s = True" by auto then show ?thesis proof (cases "\e. fetch_instruction (delayed_pool_write s) = Inl e") case True then have f3: "\e. fetch_instruction (delayed_pool_write s) = Inl e" by auto then have f4: "\ (\v. fetch_instruction (delayed_pool_write s) = Inr v)" using fetch_instr_result_3 by auto then show ?thesis using a1 f2 f3 raise_trap_result empty_delayed_pool_write_privilege raise_trap_privilege apply (simp add: execute_instruction_def) apply (simp add: exec_gets return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: case_prod_unfold) by (blast intro: empty_delayed_pool_write_privilege raise_trap_privilege) next case False then have f5: "\v. fetch_instruction (delayed_pool_write s) = Inr v" using fetch_instr_result_1 by blast then have f6: "\v. fetch_instruction (delayed_pool_write s) = Inr v \ \ (\e. ((decode_instruction v)::(Exception list + instruction)) = Inl e)" using a1 f2 dispatch_fail by blast then have f7: "\v. fetch_instruction (delayed_pool_write s) = Inr v \ (\v1. ((decode_instruction v)::(Exception list + instruction)) = Inr v1)" using decode_instr_result_4 by auto then show ?thesis proof (cases "annul_val (delayed_pool_write s)") case True then show ?thesis using a1 f2 f7 apply (simp add: execute_instruction_def) apply (simp add: exec_gets return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (auto intro: empty_delayed_pool_write_privilege set_annul_privilege write_cpu_npc_privilege write_cpu_pc_privilege) next case False then show ?thesis using a1 f2 f7 apply (simp add: execute_instruction_def) apply (simp add: exec_gets return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: simpler_modify_def return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (auto intro: empty_delayed_pool_write_privilege dispath_instr_privilege execute_instr_sub1_privilege) qed qed next case False then show ?thesis using a1 apply (simp add: execute_instruction_def) by (simp add: simpler_gets_def bind_def h1_def h2_def Let_def return_def) qed (*********************************************************************) section \Single step non-interference property.\ (*********************************************************************) -definition user_accessible:: "('a::len0) sparc_state \ phys_address \ bool" where +definition user_accessible:: "('a::len) sparc_state \ phys_address \ bool" where "user_accessible s pa \ \va p. (virt_to_phys va (mmu s) (mem s)) = Some p \ mmu_readable (get_acc_flag (snd p)) 10 \ (fst p) = pa" \ \Passing \asi = 8\ is the same.\ lemma user_accessible_8: assumes a1: "mmu_readable (get_acc_flag (snd p)) 8" shows "mmu_readable (get_acc_flag (snd p)) 10" using a1 by (simp add: mmu_readable_def) definition mem_equal:: "('a) sparc_state \ ('a) sparc_state \ phys_address \ bool" where "mem_equal s1 s2 pa \ (mem s1) 8 (pa AND 68719476732) = (mem s2) 8 (pa AND 68719476732) \ (mem s1) 8 ((pa AND 68719476732) + 1) = (mem s2) 8 ((pa AND 68719476732) + 1) \ (mem s1) 8 ((pa AND 68719476732) + 2) = (mem s2) 8 ((pa AND 68719476732) + 2) \ (mem s1) 8 ((pa AND 68719476732) + 3) = (mem s2) 8 ((pa AND 68719476732) + 3) \ (mem s1) 9 (pa AND 68719476732) = (mem s2) 9 (pa AND 68719476732) \ (mem s1) 9 ((pa AND 68719476732) + 1) = (mem s2) 9 ((pa AND 68719476732) + 1) \ (mem s1) 9 ((pa AND 68719476732) + 2) = (mem s2) 9 ((pa AND 68719476732) + 2) \ (mem s1) 9 ((pa AND 68719476732) + 3) = (mem s2) 9 ((pa AND 68719476732) + 3) \ (mem s1) 10 (pa AND 68719476732) = (mem s2) 10 (pa AND 68719476732) \ (mem s1) 10 ((pa AND 68719476732) + 1) = (mem s2) 10 ((pa AND 68719476732) + 1) \ (mem s1) 10 ((pa AND 68719476732) + 2) = (mem s2) 10 ((pa AND 68719476732) + 2) \ (mem s1) 10 ((pa AND 68719476732) + 3) = (mem s2) 10 ((pa AND 68719476732) + 3) \ (mem s1) 11 (pa AND 68719476732) = (mem s2) 11 (pa AND 68719476732) \ (mem s1) 11 ((pa AND 68719476732) + 1) = (mem s2) 11 ((pa AND 68719476732) + 1) \ (mem s1) 11 ((pa AND 68719476732) + 2) = (mem s2) 11 ((pa AND 68719476732) + 2) \ (mem s1) 11 ((pa AND 68719476732) + 3) = (mem s2) 11 ((pa AND 68719476732) + 3)" text \\low_equal\ defines the equivalence relation over two sparc states that is an analogy to the \=\<^sub>L\ relation over memory contexts in the traditional non-interference theorem.\ -definition low_equal:: "('a::len0) sparc_state \ ('a) sparc_state \ bool" where +definition low_equal:: "('a::len) sparc_state \ ('a) sparc_state \ bool" where "low_equal s1 s2 \ (cpu_reg s1) = (cpu_reg s2) \ (user_reg s1) = (user_reg s2) \ (sys_reg s1) = (sys_reg s2) \ (\va. (virt_to_phys va (mmu s1) (mem s1)) = (virt_to_phys va (mmu s2) (mem s2))) \ (\pa. (user_accessible s1 pa) \ mem_equal s1 s2 pa) \ (mmu s1) = (mmu s2) \ (state_var s1) = (state_var s2) \ (traps s1) = (traps s2) \ (undef s1) = (undef s2) " lemma low_equal_com: "low_equal s1 s2 \ low_equal s2 s1" apply (simp add: low_equal_def) apply (simp add: mem_equal_def user_accessible_def) by metis lemma non_exe_mode_equal: "exe_mode_val s = False \ get_trap_set s = {} \ Some t = NEXT s \ t = s" apply (simp add: NEXT_def execute_instruction_def) apply auto by (simp add: simpler_gets_def bind_def h1_def h2_def Let_def return_def) lemma exe_mode_low_equal: assumes a1: "low_equal s1 s2" shows " exe_mode_val s1 = exe_mode_val s2" using a1 apply (simp add: low_equal_def) by (simp add: exe_mode_val_def) lemma mem_val_mod_state: "mem_val_alt asi a s = mem_val_alt asi a (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)" apply (simp add: mem_val_alt_def) by (simp add: Let_def) lemma mem_val_w32_mod_state: "mem_val_w32 asi a s = mem_val_w32 asi a (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)" apply (simp add: mem_val_w32_def) apply (simp add: Let_def) by (metis mem_val_mod_state) lemma load_word_mem_mod_state: "load_word_mem s addr asi = load_word_mem (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr asi" apply (simp add: load_word_mem_def) apply (case_tac "virt_to_phys addr (mmu s) (mem s) = None") apply auto by (auto simp add: mem_val_w32_mod_state) lemma load_word_mem2_mod_state: "fst (case load_word_mem s addr asi of None \ (None, s) | Some w \ (Some w, add_data_cache s addr w 15)) = fst (case load_word_mem (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr asi of None \ (None, (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)) | Some w \ (Some w, add_data_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr w 15))" proof (cases "load_word_mem s addr asi = None") case True then have "load_word_mem s addr asi = None \ load_word_mem (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr asi = None" using load_word_mem_mod_state by metis then show ?thesis by auto next case False then have "\w. load_word_mem s addr asi = Some w" by auto then have "\w. load_word_mem s addr asi = Some w \ load_word_mem (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr asi = Some w" using load_word_mem_mod_state by metis then show ?thesis by auto qed lemma load_word_mem3_mod_state: "fst (case load_word_mem s addr asi of None \ (None, s) | Some w \ (Some w, add_instr_cache s addr w 15)) = fst (case load_word_mem (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr asi of None \ (None, (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)) | Some w \ (Some w, add_instr_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr w 15))" proof (cases "load_word_mem s addr asi = None") case True then have "load_word_mem s addr asi = None \ load_word_mem (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr asi = None" using load_word_mem_mod_state by metis then show ?thesis by auto next case False then have "\w. load_word_mem s addr asi = Some w" by auto then have "\w. load_word_mem s addr asi = Some w \ load_word_mem (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr asi = Some w" using load_word_mem_mod_state by metis then show ?thesis by auto qed lemma read_dcache_mod_state: "read_data_cache s addr = read_data_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr" apply (simp add: read_data_cache_def) by (simp add: dcache_val_def) lemma read_dcache2_mod_state: "fst (case read_data_cache s addr of None \ (None, s) | Some w \ (Some w, s)) = fst (case read_data_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr of None \ (None, (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)) | Some w \ (Some w, (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)))" proof (cases "read_data_cache s addr = None") case True then have "read_data_cache s addr = None \ read_data_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr = None" using read_dcache_mod_state by metis then show ?thesis by auto next case False then have "\w. read_data_cache s addr = Some w" by auto then have "\w. read_data_cache s addr = Some w \ read_data_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr = Some w" using read_dcache_mod_state by metis then show ?thesis by auto qed lemma read_icache_mod_state: "read_instr_cache s addr = read_instr_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr" apply (simp add: read_instr_cache_def) by (simp add: icache_val_def) lemma read_icache2_mod_state: "fst (case read_instr_cache s addr of None \ (None, s) | Some w \ (Some w, s)) = fst (case read_instr_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr of None \ (None, (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)) | Some w \ (Some w, (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)))" proof (cases "read_instr_cache s addr = None") case True then have "read_instr_cache s addr = None \ read_instr_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr = None" using read_icache_mod_state by metis then show ?thesis by auto next case False then have "\w. read_instr_cache s addr = Some w" by auto then have "\w. read_instr_cache s addr = Some w \ read_instr_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr = Some w" using read_icache_mod_state by metis then show ?thesis by auto qed lemma mem_read_mod_state: "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\))" apply (simp add: memory_read_def) apply (case_tac "uint asi = 1") apply (simp add: Let_def) apply (metis load_word_mem_mod_state option.distinct(1)) apply (case_tac "uint asi = 2") apply (simp add: Let_def) apply (simp add: sys_reg_val_def) apply (case_tac "uint asi \ {8,9}") apply (simp add: Let_def) apply (simp add: load_word_mem3_mod_state) apply (simp add: load_word_mem_mod_state) apply (case_tac "uint asi \ {10,11}") apply (simp add: Let_def) apply (simp add: load_word_mem2_mod_state) apply (simp add: load_word_mem_mod_state) apply (case_tac "uint asi = 13") apply (simp add: Let_def) apply (simp add: read_icache2_mod_state) apply (case_tac "uint asi = 15") apply (simp add: Let_def) apply (simp add: read_dcache2_mod_state) apply (case_tac "uint asi = 25") apply (simp add: Let_def) apply (case_tac "uint asi = 28") apply (simp add: Let_def) apply (simp add: mem_val_w32_mod_state) by (simp add: Let_def) lemma insert_trap_mem: "fst (memory_read asi addr s) = fst (memory_read asi addr (s\traps := new_traps\))" proof - have "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := (cpu_reg s), user_reg := (user_reg s), dwrite := (dwrite s), state_var := (state_var s), traps := new_traps, undef := (undef s)\))" using mem_read_mod_state by blast then show ?thesis by auto qed lemma cpu_reg_mod_mem: "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := new_cpu_reg\))" proof - have "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := new_cpu_reg, user_reg := (user_reg s), dwrite := (dwrite s), state_var := (state_var s), traps := (traps s), undef := (undef s)\))" using mem_read_mod_state by blast then show ?thesis by auto qed lemma user_reg_mod_mem: "fst (memory_read asi addr s) = fst (memory_read asi addr (s\user_reg := new_user_reg\))" proof - have "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := (cpu_reg s), user_reg := new_user_reg, dwrite := (dwrite s), state_var := (state_var s), traps := (traps s), undef := (undef s)\))" using mem_read_mod_state by blast then show ?thesis by auto qed lemma annul_mem: "fst (memory_read asi addr s) = fst (memory_read asi addr (s\state_var := new_state_var, cpu_reg := new_cpu_reg\))" proof - have "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := new_cpu_reg, user_reg := (user_reg s), dwrite := (dwrite s), state_var := new_state_var, traps := (traps s), undef := (undef s)\))" using mem_read_mod_state by blast then have "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := new_cpu_reg, state_var := new_state_var\))" by auto then show ?thesis by (metis Sparc_State.sparc_state.surjective Sparc_State.sparc_state.update_convs(1) Sparc_State.sparc_state.update_convs(8)) qed lemma state_var_mod_mem: "fst (memory_read asi addr s) = fst (memory_read asi addr (s\state_var := new_state_var\))" proof - have "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := (cpu_reg s), user_reg := (user_reg s), dwrite := (dwrite s), state_var := new_state_var, traps := (traps s), undef := (undef s)\))" using mem_read_mod_state by blast then show ?thesis by auto qed lemma mod_state_low_equal: "low_equal s1 s2 \ t1 = (s1\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) \ t2 = (s2\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) \ low_equal t1 t2" apply (simp add: low_equal_def) apply clarsimp apply (simp add: mem_equal_def) by (simp add: user_accessible_def) lemma user_reg_state_mod_low_equal: assumes a1: "low_equal s1 s2 \ t1 = (s1\user_reg := new_user_reg\) \ t2 = (s2\user_reg := new_user_reg\)" shows "low_equal t1 t2" proof - have "low_equal s1 s2 \ t1 = (s1\cpu_reg := (cpu_reg s1), user_reg := new_user_reg, dwrite := (dwrite s1), state_var := (state_var s1), traps := (traps s1), undef := (undef s1)\) \ t2 = (s2\cpu_reg := (cpu_reg s2), user_reg := new_user_reg, dwrite := (dwrite s2), state_var := (state_var s2), traps := (traps s2), undef := (undef s2)\) \ low_equal t1 t2" using mod_state_low_equal apply (simp add: low_equal_def) apply (simp add: user_accessible_def mem_equal_def) by clarsimp then show ?thesis using a1 by clarsimp qed lemma mod_trap_low_equal: assumes a1: "low_equal s1 s2 \ t1 = (s1\traps := new_traps\) \ t2 = (s2\traps := new_traps\)" shows "low_equal t1 t2" proof - have "low_equal s1 s2 \ t1 = (s1\cpu_reg := (cpu_reg s1), user_reg := (user_reg s1), dwrite := (dwrite s1), state_var := (state_var s1), traps := new_traps, undef := (undef s1)\) \ t2 = (s2\cpu_reg := (cpu_reg s2), user_reg := (user_reg s2), dwrite := (dwrite s2), state_var := (state_var s2), traps := new_traps, undef := (undef s2)\) \ low_equal t1 t2" using mod_state_low_equal apply (simp add: low_equal_def) apply (simp add: user_accessible_def mem_equal_def) by clarsimp then show ?thesis using a1 by clarsimp qed lemma state_var_low_equal: "low_equal s1 s2 \ state_var s1 = state_var s2" by (simp add: low_equal_def) lemma state_var2_low_equal: assumes a1: "low_equal s1 s2 \ t1 = (s1\state_var := new_state_var\) \ t2 = (s2\state_var := new_state_var\)" shows "low_equal t1 t2" proof - have "low_equal s1 s2 \ t1 = (s1\cpu_reg := (cpu_reg s1), user_reg := (user_reg s1), dwrite := (dwrite s1), state_var := new_state_var, traps := (traps s1), undef := (undef s1)\) \ t2 = (s2\cpu_reg := (cpu_reg s2), user_reg := (user_reg s2), dwrite := (dwrite s2), state_var := new_state_var, traps := (traps s2), undef := (undef s2)\) \ low_equal t1 t2" using mod_state_low_equal apply (simp add: low_equal_def) apply (simp add: user_accessible_def mem_equal_def) by clarsimp then show ?thesis using a1 by clarsimp qed lemma traps_low_equal: "low_equal s1 s2 \ traps s1 = traps s2" by (simp add: low_equal_def) lemma s_low_equal: "low_equal s1 s2 \ (get_S (cpu_reg_val PSR s1)) = (get_S (cpu_reg_val PSR s2))" by (simp add: low_equal_def cpu_reg_val_def) lemma cpu_reg_val_low_equal: "low_equal s1 s2 \ (cpu_reg_val cr s1) = (cpu_reg_val cr s2)" by (simp add: cpu_reg_val_def low_equal_def) lemma get_curr_win_low_equal: "low_equal s1 s2 \ (fst (fst (get_curr_win () s1))) = (fst (fst (get_curr_win () s2)))" apply (simp add: low_equal_def) apply (simp add: get_curr_win_def cpu_reg_val_def get_CWP_def) by (simp add: simpler_gets_def) lemma get_curr_win2_low_equal: "low_equal s1 s2 \ t1 = (snd (fst (get_curr_win () s1))) \ t2 = (snd (fst (get_curr_win () s2))) \ low_equal t1 t2" apply (simp add: low_equal_def) apply (simp add: get_curr_win_def cpu_reg_val_def get_CWP_def) by (auto simp add: simpler_gets_def) lemma get_curr_win3_low_equal: "low_equal s1 s2 \ (traps (snd (fst (get_curr_win () s1)))) = (traps (snd (fst (get_curr_win () s2))))" using low_equal_def get_curr_win2_low_equal by blast lemma get_addr_low_equal: "low_equal s1 s2 \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) = ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word3) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) = ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word2) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word1) = ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word1)" apply (simp add: low_equal_def) apply (simp add: get_curr_win_def cpu_reg_val_def get_CWP_def) apply (simp add: simpler_gets_def get_addr_def user_reg_val_def) apply (simp add: Let_def ) apply (simp add: get_CWP_def cpu_reg_val_def get_operand2_def) by (simp add: user_reg_val_def) lemma get_addr2_low_equal: "low_equal s1 s2 \ get_addr (snd instr) (snd (fst (get_curr_win () s1))) = get_addr (snd instr) (snd (fst (get_curr_win () s2)))" apply (simp add: low_equal_def) apply (simp add: get_curr_win_def cpu_reg_val_def get_CWP_def) apply (simp add: simpler_gets_def get_addr_def user_reg_val_def) apply (simp add: Let_def ) apply (simp add: get_CWP_def cpu_reg_val_def get_operand2_def) by (simp add: user_reg_val_def) lemma sys_reg_low_equal: "low_equal s1 s2 \ sys_reg s1 = sys_reg s2" by (simp add: low_equal_def) lemma user_reg_low_equal: "low_equal s1 s2 \ user_reg s1 = user_reg s2" by (simp add: low_equal_def) lemma user_reg_val_low_equal: "low_equal s1 s2 \ user_reg_val win ur s1 = user_reg_val win ur s2" apply (simp add: user_reg_val_def) by (simp add: user_reg_low_equal) lemma get_operand2_low_equal: "low_equal s1 s2 \ get_operand2 op_list s1 = get_operand2 op_list s2" apply (simp add: get_operand2_def) apply (simp add: cpu_reg_val_low_equal) apply auto apply (simp add: user_reg_val_def) using user_reg_low_equal by fastforce lemma mem_val_mod_cache: "mem_val_alt asi a s = mem_val_alt asi a (s\cache := new_cache\)" apply (simp add: mem_val_alt_def) by (simp add: Let_def) lemma mem_val_w32_mod_cache: "mem_val_w32 asi a s = mem_val_w32 asi a (s\cache := new_cache\)" apply (simp add: mem_val_w32_def) apply (simp add: Let_def) by (metis mem_val_mod_cache) lemma load_word_mem_mod_cache: "load_word_mem s addr asi = load_word_mem (s\cache := new_cache\) addr asi" apply (simp add: load_word_mem_def) apply (case_tac "virt_to_phys addr (mmu s) (mem s) = None") apply auto by (simp add: mem_val_w32_mod_cache) lemma memory_read_8_mod_cache: "fst (memory_read 8 addr s) = fst (memory_read 8 addr (s\cache := new_cache\))" apply (simp add: memory_read_def) apply (case_tac "sys_reg s CCR AND 1 \ 0") apply auto apply (simp add: option.case_eq_if load_word_mem_mod_cache) apply (auto intro: load_word_mem_mod_cache) apply (metis load_word_mem_mod_cache option.distinct(1)) by (metis load_word_mem_mod_cache option.distinct(1)) lemma memory_read_10_mod_cache: "fst (memory_read 10 addr s) = fst (memory_read 10 addr (s\cache := new_cache\))" apply (simp add: memory_read_def) apply (case_tac "sys_reg s CCR AND 1 \ 0") apply auto apply (simp add: option.case_eq_if load_word_mem_mod_cache) apply (auto intro: load_word_mem_mod_cache) apply (metis load_word_mem_mod_cache option.distinct(1)) by (metis load_word_mem_mod_cache option.distinct(1)) lemma mem_equal_mod_cache: "mem_equal s1 s2 pa \ mem_equal (s1\cache := new_cache1\) (s2\cache := new_cache2\) pa" by (simp add: mem_equal_def) lemma user_accessible_mod_cache: "user_accessible (s\cache := new_cache\) pa = user_accessible s pa" by (simp add: user_accessible_def) lemma mem_equal_mod_user_reg: "mem_equal s1 s2 pa \ mem_equal (s1\user_reg := new_user_reg1\) (s2\user_reg := user_reg2\) pa" by (simp add: mem_equal_def) lemma user_accessible_mod_user_reg: "user_accessible (s\user_reg := new_user_reg\) pa = user_accessible s pa" by (simp add: user_accessible_def) lemma mem_equal_mod_cpu_reg: "mem_equal s1 s2 pa \ mem_equal (s1\cpu_reg := new_cpu1\) (s2\cpu_reg := cpu_reg2\) pa" by (simp add: mem_equal_def) lemma user_accessible_mod_cpu_reg: "user_accessible (s\cpu_reg := new_cpu_reg\) pa = user_accessible s pa" by (simp add: user_accessible_def) lemma mem_equal_mod_trap: "mem_equal s1 s2 pa \ mem_equal (s1\traps := new_traps1\) (s2\traps := traps2\) pa" by (simp add: mem_equal_def) lemma user_accessible_mod_trap: "user_accessible (s\traps := new_traps\) pa = user_accessible s pa" by (simp add: user_accessible_def) lemma mem_equal_annul: "mem_equal s1 s2 pa \ mem_equal (s1\state_var := new_state_var, cpu_reg := new_cpu_reg\) (s2\state_var := new_state_var2, cpu_reg := new_cpu_reg2\) pa" by (simp add: mem_equal_def) lemma user_accessible_annul: "user_accessible (s\state_var := new_state_var, cpu_reg := new_cpu_reg\) pa = user_accessible s pa" by (simp add: user_accessible_def) lemma mem_val_alt_10_mem_equal_0: "mem_equal s1 s2 pa \ mem_val_alt 10 (pa AND 68719476732) s1 = mem_val_alt 10 (pa AND 68719476732) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_mem_equal_1: "mem_equal s1 s2 pa \ mem_val_alt 10 ((pa AND 68719476732) + 1) s1 = mem_val_alt 10 ((pa AND 68719476732) + 1) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_mem_equal_2: "mem_equal s1 s2 pa \ mem_val_alt 10 ((pa AND 68719476732) + 2) s1 = mem_val_alt 10 ((pa AND 68719476732) + 2) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_mem_equal_3: "mem_equal s1 s2 pa \ mem_val_alt 10 ((pa AND 68719476732) + 3) s1 = mem_val_alt 10 ((pa AND 68719476732) + 3) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_mem_equal: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 10 (pa AND 68719476732) s1 = mem_val_alt 10 (pa AND 68719476732) s2 \ mem_val_alt 10 ((pa AND 68719476732) + 1) s1 = mem_val_alt 10 ((pa AND 68719476732) + 1) s2 \ mem_val_alt 10 ((pa AND 68719476732) + 2) s1 = mem_val_alt 10 ((pa AND 68719476732) + 2) s2 \ mem_val_alt 10 ((pa AND 68719476732) + 3) s1 = mem_val_alt 10 ((pa AND 68719476732) + 3) s2" using mem_val_alt_10_mem_equal_0 mem_val_alt_10_mem_equal_1 mem_val_alt_10_mem_equal_2 mem_val_alt_10_mem_equal_3 a1 by blast lemma mem_val_w32_10_mem_equal: assumes a1: "mem_equal s1 s2 a" shows "mem_val_w32 10 a s1 = mem_val_w32 10 a s2" apply (simp add: mem_val_w32_def) apply (simp add: Let_def) using mem_val_alt_10_mem_equal a1 apply auto apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce by fastforce lemma mem_val_alt_8_mem_equal_0: "mem_equal s1 s2 pa \ mem_val_alt 8 (pa AND 68719476732) s1 = mem_val_alt 8 (pa AND 68719476732) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_mem_equal_1: "mem_equal s1 s2 pa \ mem_val_alt 8 ((pa AND 68719476732) + 1) s1 = mem_val_alt 8 ((pa AND 68719476732) + 1) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_mem_equal_2: "mem_equal s1 s2 pa \ mem_val_alt 8 ((pa AND 68719476732) + 2) s1 = mem_val_alt 8 ((pa AND 68719476732) + 2) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_mem_equal_3: "mem_equal s1 s2 pa \ mem_val_alt 8 ((pa AND 68719476732) + 3) s1 = mem_val_alt 8 ((pa AND 68719476732) + 3) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_mem_equal: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 8 (pa AND 68719476732) s1 = mem_val_alt 8 (pa AND 68719476732) s2 \ mem_val_alt 8 ((pa AND 68719476732) + 1) s1 = mem_val_alt 8 ((pa AND 68719476732) + 1) s2 \ mem_val_alt 8 ((pa AND 68719476732) + 2) s1 = mem_val_alt 8 ((pa AND 68719476732) + 2) s2 \ mem_val_alt 8 ((pa AND 68719476732) + 3) s1 = mem_val_alt 8 ((pa AND 68719476732) + 3) s2" using mem_val_alt_8_mem_equal_0 mem_val_alt_8_mem_equal_1 mem_val_alt_8_mem_equal_2 mem_val_alt_8_mem_equal_3 a1 by blast lemma mem_val_w32_8_mem_equal: assumes a1: "mem_equal s1 s2 a" shows "mem_val_w32 8 a s1 = mem_val_w32 8 a s2" apply (simp add: mem_val_w32_def) apply (simp add: Let_def) using mem_val_alt_8_mem_equal a1 apply auto apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce by fastforce lemma load_word_mem_10_low_equal: assumes a1: "low_equal s1 s2" shows "load_word_mem s1 address 10 = load_word_mem s2 address 10" using a1 apply (simp add: low_equal_def load_word_mem_def) apply clarsimp apply (case_tac "virt_to_phys address (mmu s2) (mem s2) = None") apply auto apply (simp add: user_accessible_def) using mem_val_w32_10_mem_equal apply blast apply (simp add: user_accessible_def) using mem_val_w32_10_mem_equal by blast lemma load_word_mem_8_low_equal: assumes a1: "low_equal s1 s2" shows "load_word_mem s1 address 8 = load_word_mem s2 address 8" using a1 apply (simp add: low_equal_def load_word_mem_def) apply clarsimp apply (case_tac "virt_to_phys address (mmu s2) (mem s2) = None") apply auto apply (simp add: user_accessible_def) using mem_val_w32_8_mem_equal user_accessible_8 apply fastforce apply (simp add: user_accessible_def) using mem_val_w32_8_mem_equal user_accessible_8 by fastforce lemma mem_read_low_equal: assumes a1: "low_equal s1 s2 \ asi \ {8,10}" shows "fst (memory_read asi address s1) = fst (memory_read asi address s2)" proof (cases "asi = 8") case True then show ?thesis using a1 apply (simp add: low_equal_def) apply (simp add: memory_read_def) using a1 load_word_mem_8_low_equal apply auto apply (simp add: option.case_eq_if) by (simp add: option.case_eq_if) next case False then have "asi = 10" using a1 by auto then show ?thesis using a1 apply (simp add: low_equal_def) apply (simp add: memory_read_def) using a1 load_word_mem_10_low_equal apply auto apply (simp add: option.case_eq_if) by (simp add: option.case_eq_if) qed lemma read_mem_pc_low_equal: assumes a1: "low_equal s1 s2" shows "fst (memory_read 8 (cpu_reg_val PC s1) s1) = fst (memory_read 8 (cpu_reg_val PC s2) s2)" proof - have f2: "cpu_reg_val PC s1 = cpu_reg_val PC s2" using a1 by (simp add: low_equal_def cpu_reg_val_def) then show ?thesis using a1 f2 mem_read_low_equal by auto qed lemma dcache_mod_low_equal: assumes a1: "low_equal s1 s2 \ t1 = dcache_mod c v s1 \ t2 = dcache_mod c v s2" shows "low_equal t1 t2" using a1 apply (simp add: low_equal_def) apply (simp add: dcache_mod_def) apply auto apply (simp add: user_accessible_mod_cache mem_equal_mod_cache) by (simp add: user_accessible_mod_cache mem_equal_mod_cache) lemma add_data_cache_low_equal: assumes a1: "low_equal s1 s2 \ t1 = add_data_cache s1 address w bm \ t2 = add_data_cache s2 address w bm" shows "low_equal t1 t2" using a1 apply (simp add: add_data_cache_def) apply (case_tac "bm AND 8 >> 3 = 1") apply auto apply (case_tac "bm AND 4 >> 2 = 1") apply auto apply (case_tac "bm AND 2 >> Suc 0 = 1") apply auto apply (case_tac "bm AND 1 = 1") apply auto apply (meson dcache_mod_low_equal) apply (meson dcache_mod_low_equal) apply (case_tac "bm AND 1 = 1") apply auto apply (meson dcache_mod_low_equal) apply (meson dcache_mod_low_equal) apply (case_tac "bm AND 2 >> Suc 0 = 1") apply auto apply (case_tac "bm AND 1 = 1") apply auto apply (meson dcache_mod_low_equal) apply (meson dcache_mod_low_equal) apply (case_tac "bm AND 1 = 1") apply auto apply (meson dcache_mod_low_equal) apply (meson dcache_mod_low_equal) apply (case_tac "bm AND 4 >> 2 = 1") apply auto apply (case_tac "bm AND 2 >> Suc 0 = 1") apply auto apply (case_tac "bm AND 1 = 1") apply auto apply (meson dcache_mod_low_equal) apply (meson dcache_mod_low_equal) apply (case_tac "bm AND 1 = 1") apply auto apply (meson dcache_mod_low_equal) apply (meson dcache_mod_low_equal) apply (case_tac "bm AND 2 >> Suc 0 = 1") apply auto apply (case_tac "bm AND 1 = 1") apply auto apply (meson dcache_mod_low_equal) apply (meson dcache_mod_low_equal) by (meson dcache_mod_low_equal) lemma mem_read2_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (memory_read (10::word8) address s1) \ t2 = snd (memory_read (10::word8) address s2)" shows "low_equal t1 t2" using a1 apply (simp add: memory_read_def) using a1 apply (auto simp add: sys_reg_low_equal) using a1 apply (simp add: load_word_mem_10_low_equal) by (metis (no_types, lifting) add_data_cache_low_equal option.case_eq_if snd_conv) lemma mem_read_delayed_write_low_equal: assumes a1: "low_equal s1 s2 \ get_delayed_pool s1 = [] \ get_delayed_pool s2 = []" shows "fst (memory_read 8 (cpu_reg_val PC (delayed_pool_write s1)) (delayed_pool_write s1)) = fst (memory_read 8 (cpu_reg_val PC (delayed_pool_write s2)) (delayed_pool_write s2))" using a1 apply (simp add: delayed_pool_write_def) apply (simp add: Let_def) apply (simp add: get_delayed_write_def) by (simp add: read_mem_pc_low_equal) lemma global_reg_mod_low_equal: assumes a1: "low_equal s1 s2\ t1 = (global_reg_mod w n rd s1) \ t2 = (global_reg_mod w n rd s2)" shows "low_equal t1 t2" using a1 apply (induction n arbitrary: s1 s2) apply clarsimp apply auto apply (simp add: Let_def) apply (simp add: user_reg_low_equal) using user_reg_state_mod_low_equal by blast lemma out_reg_mod_low_equal: assumes a1: "low_equal s1 s2\ t1 = (out_reg_mod w curr_win rd s1) \ t2 = (out_reg_mod w curr_win rd s2)" shows "low_equal t1 t2" using a1 apply (simp add: out_reg_mod_def Let_def) apply auto apply (simp add: user_reg_low_equal) using user_reg_state_mod_low_equal apply fastforce apply (simp add: user_reg_low_equal) using user_reg_state_mod_low_equal by blast lemma in_reg_mod_low_equal: assumes a1: "low_equal s1 s2\ t1 = (in_reg_mod w curr_win rd s1) \ t2 = (in_reg_mod w curr_win rd s2)" shows "low_equal t1 t2" using a1 apply (simp add: in_reg_mod_def Let_def) apply auto apply (simp add: user_reg_low_equal) using user_reg_state_mod_low_equal apply fastforce apply (simp add: user_reg_low_equal) using user_reg_state_mod_low_equal by blast lemma user_reg_mod_low_equal: assumes a1: "low_equal s1 s2 \ t1 = user_reg_mod w curr_win rd s1 \ t2 = user_reg_mod w curr_win rd s2" shows "low_equal t1 t2" proof (cases "rd = 0") case True then show ?thesis using a1 by (simp add: user_reg_mod_def) next case False then have f1: "rd \ 0" by auto then show ?thesis proof (cases "0 < rd \ rd < 8") case True then show ?thesis using a1 f1 apply (simp add: user_reg_mod_def) using global_reg_mod_low_equal by blast next case False then have f2: "\ (0 < rd \ rd < 8)" by auto then show ?thesis proof (cases "7 < rd \ rd < 16") case True then show ?thesis using a1 f1 f2 apply (simp add: user_reg_mod_def) by (auto intro: out_reg_mod_low_equal) next case False then have f3: "\ (7 < rd \ rd < 16)" by auto then show ?thesis proof (cases "15 < rd \ rd < 24") case True then show ?thesis using a1 f1 f2 f3 apply (simp add: user_reg_mod_def) apply (simp add: low_equal_def) apply clarsimp by (simp add: user_accessible_mod_user_reg mem_equal_mod_user_reg) next case False then show ?thesis using a1 f1 f2 f3 apply (simp add: user_reg_mod_def) by (auto intro: in_reg_mod_low_equal) qed qed qed qed lemma virt_to_phys_low_equal: "low_equal s1 s2 \ virt_to_phys addr (mmu s1) (mem s1) = virt_to_phys addr (mmu s2) (mem s2)" by (auto simp add: low_equal_def) lemma write_reg_low_equal: assumes a1: "low_equal s1 s2 \ t1 = (snd (fst (write_reg w curr_win rd s1))) \ t2 = (snd (fst (write_reg w curr_win rd s2)))" shows "low_equal t1 t2" using a1 apply (simp add: write_reg_def) apply (simp add: simpler_modify_def) by (auto intro: user_reg_mod_low_equal) lemma write_cpu_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (write_cpu w cr s1)) \ t2 = (snd (fst (write_cpu w cr s2)))" shows "low_equal t1 t2" using a1 apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: cpu_reg_mod_def) apply (simp add: low_equal_def) using user_accessible_mod_cpu_reg mem_equal_mod_cpu_reg by metis lemma cpu_reg_mod_low_equal: assumes a1: "low_equal s1 s2 \ t1 = cpu_reg_mod w cr s1 \ t2 = cpu_reg_mod w cr s2" shows "low_equal t1 t2" using a1 apply (simp add: cpu_reg_mod_def) apply (simp add: low_equal_def) using user_accessible_mod_cpu_reg mem_equal_mod_cpu_reg by metis lemma load_sub2_low_equal: assumes a1: "low_equal s1 s2 \ t1 = (snd (fst (load_sub2 address 10 rd curr_win w s1))) \ t2 = (snd (fst (load_sub2 address 10 rd curr_win w s2)))" shows "low_equal t1 t2" proof (cases "fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) = None") case True then have f0: "fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) = None" by auto have f1: "low_equal (snd (fst (write_reg w curr_win (rd AND 30) s1))) (snd (fst (write_reg w curr_win (rd AND 30) s2)))" using a1 by (auto intro: write_reg_low_equal) then have "fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) = None \ fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) = fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s2))))" using f0 by (blast intro: mem_read_low_equal) then have "fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) = None \ fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s2)))) = None" by auto then show ?thesis using a1 apply (simp add: load_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) using f1 apply (simp add: traps_low_equal) using f1 by (auto intro: mod_trap_low_equal) next case False then have f2: "fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) \ None" by auto have f3: "low_equal (snd (fst (write_reg w curr_win (rd AND 30) s1))) (snd (fst (write_reg w curr_win (rd AND 30) s2)))" using a1 by (auto intro: write_reg_low_equal) then have f4: "fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) = fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s2))))" using f2 by (blast intro: mem_read_low_equal) then have "fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) \ None \ fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s2)))) \ None" using f2 by auto then show ?thesis using a1 apply (simp add: load_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply clarsimp apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) using f4 apply clarsimp using f3 by (auto intro: mem_read2_low_equal write_reg_low_equal) qed lemma load_sub3_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (load_sub3 instr curr_win rd (10::word8) address s1)) \ t2 = snd (fst (load_sub3 instr curr_win rd (10::word8) address s2))" shows "low_equal t1 t2" proof (cases "fst (memory_read 10 address s1) = None") case True then have "fst (memory_read 10 address s1) = None \ fst (memory_read 10 address s2) = None" using a1 by (auto simp add: mem_read_low_equal) then show ?thesis using a1 apply (simp add: load_sub3_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply (auto simp add: traps_low_equal) by (auto intro: mod_trap_low_equal) next case False then have f1: "fst (memory_read 10 address s1) \ None \ fst (memory_read 10 address s2) \ None" using a1 by (auto simp add: mem_read_low_equal) then show ?thesis proof (cases "rd \ 0 \ (fst instr = load_store_type LD \ fst instr = load_store_type LDA \ fst instr = load_store_type LDUH \ fst instr = load_store_type LDSB \ fst instr = load_store_type LDUB \ fst instr = load_store_type LDUBA \ fst instr = load_store_type LDSH \ fst instr = load_store_type LDSHA \ fst instr = load_store_type LDUHA \ fst instr = load_store_type LDSBA)") case True then show ?thesis using a1 f1 apply (simp add: load_sub3_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply clarsimp apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) apply (simp add: mem_read_low_equal) by (meson mem_read2_low_equal write_reg_low_equal) next case False then show ?thesis using a1 f1 apply (simp add: load_sub3_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply clarsimp apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) apply (simp add: mem_read_low_equal) by (meson load_sub2_low_equal mem_read2_low_equal) qed qed lemma ld_asi_user: "(fst instr = load_store_type LDSB \ fst instr = load_store_type LDUB \ fst instr = load_store_type LDUH \ fst instr = load_store_type LD \ fst instr = load_store_type LDD) \ ld_asi instr 0 = 10" apply (simp add: ld_asi_def) by auto lemma load_sub1_low_equal: assumes a1: "low_equal s1 s2 \ (fst instr = load_store_type LDSB \ fst instr = load_store_type LDUB \ fst instr = load_store_type LDUH \ fst instr = load_store_type LD \ fst instr = load_store_type LDD) \ t1 = snd (fst (load_sub1 instr rd 0 s1)) \ t2 = snd (fst (load_sub1 instr rd 0 s2))" shows "low_equal t1 t2" proof (cases "(fst instr = load_store_type LDD \ fst instr = load_store_type LDDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) \ 0 \ (fst instr = load_store_type LD \ fst instr = load_store_type LDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) \ 0 \ (fst instr = load_store_type LDUH \ fst instr = load_store_type LDUHA \ fst instr = load_store_type LDSH \ fst instr = load_store_type LDSHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word1) \ 0") case True then have "((fst instr = load_store_type LDD \ fst instr = load_store_type LDDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) \ 0 \ (fst instr = load_store_type LD \ fst instr = load_store_type LDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) \ 0 \ (fst instr = load_store_type LDUH \ fst instr = load_store_type LDUHA \ fst instr = load_store_type LDSH \ fst instr = load_store_type LDSHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word1) \ 0) \ ((fst instr = load_store_type LDD \ fst instr = load_store_type LDDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word3) \ 0 \ (fst instr = load_store_type LD \ fst instr = load_store_type LDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word2) \ 0 \ (fst instr = load_store_type LDUH \ fst instr = load_store_type LDUHA \ fst instr = load_store_type LDSH \ fst instr = load_store_type LDSHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word1) \ 0)" by (metis (mono_tags, lifting) assms get_addr_low_equal) then show ?thesis using a1 apply (simp add: load_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: get_curr_win3_low_equal) by (auto intro: get_curr_win2_low_equal mod_trap_low_equal) next case False then have f1: "\ ((fst instr = load_store_type LDD \ fst instr = load_store_type LDDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) \ 0 \ (fst instr = load_store_type LD \ fst instr = load_store_type LDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) \ 0 \ (fst instr = load_store_type LDUH \ fst instr = load_store_type LDUHA \ fst instr = load_store_type LDSH \ fst instr = load_store_type LDSHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word1) \ 0) \ \ ((fst instr = load_store_type LDD \ fst instr = load_store_type LDDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word3) \ 0 \ (fst instr = load_store_type LD \ fst instr = load_store_type LDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word2) \ 0 \ (fst instr = load_store_type LDUH \ fst instr = load_store_type LDUHA \ fst instr = load_store_type LDSH \ fst instr = load_store_type LDSHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word1) \ 0)" by (metis assms get_addr_low_equal) show ?thesis proof - have "low_equal s1 s2 \ low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using get_curr_win2_low_equal by auto then have f2: "low_equal s1 s2 \ low_equal (snd (fst (load_sub3 instr (fst (fst (get_curr_win () s2))) rd 10 (get_addr (snd instr) (snd (fst (get_curr_win () s2)))) (snd (fst (get_curr_win () s1)))))) (snd (fst (load_sub3 instr (fst (fst (get_curr_win () s2))) rd 10 (get_addr (snd instr) (snd (fst (get_curr_win () s2)))) (snd (fst (get_curr_win () s2))))))" using load_sub3_low_equal by blast show ?thesis using a1 unfolding load_sub1_def simpler_gets_def bind_def h1_def h2_def Let_def case_prod_unfold using f1 f2 apply clarsimp by (simp add: get_addr2_low_equal get_curr_win_low_equal ld_asi_user) qed qed lemma load_instr_low_equal: assumes a1: "low_equal s1 s2 \ (fst instr = load_store_type LDSB \ fst instr = load_store_type LDUB \ fst instr = load_store_type LDUBA \ fst instr = load_store_type LDUH \ fst instr = load_store_type LD \ fst instr = load_store_type LDA \ fst instr = load_store_type LDD) \ ((ucast (get_S (cpu_reg_val PSR s1)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s2)))::word1) = 0 \ t1 = snd (fst (load_instr instr s1)) \ t2 = snd (fst (load_instr instr s2))" shows "low_equal t1 t2" proof - have "get_S (cpu_reg_val PSR s1) = 0 \ get_S (cpu_reg_val PSR s2) = 0" using a1 by (simp add: ucast_id) then show ?thesis using a1 apply (simp add: load_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) apply clarsimp apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply (simp add: traps_low_equal) by (auto intro: mod_trap_low_equal load_sub1_low_equal) qed lemma st_data0_low_equal: "low_equal s1 s2 \ st_data0 instr curr_win rd addr s1 = st_data0 instr curr_win rd addr s2" apply (simp add: st_data0_def) by (simp add: user_reg_val_def low_equal_def) lemma store_word_mem_low_equal_none: "low_equal s1 s2 \ store_word_mem (add_data_cache s1 addr data bm) addr data bm 10 = None \ store_word_mem (add_data_cache s2 addr data bm) addr data bm 10 = None" apply (simp add: store_word_mem_def) proof - assume a1: "low_equal s1 s2" assume a2: "(case virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) of None \ None | Some pair \ if mmu_writable (get_acc_flag (snd pair)) 10 then Some (mem_mod_w32 10 (fst pair) bm data (add_data_cache s1 addr data bm)) else None) = None" have f3: "(if mmu_writable (get_acc_flag (snd (v1_2 (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (v1_2 (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s2 addr data bm)) else None) = (case Some (v1_2 (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) of None \ if mmu_writable (get_acc_flag (snd (v1_2 (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (v1_2 (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s1 addr data bm)) else None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None)" by auto obtain pp :: "(word36 \ word8) option \ word36 \ word8" where f4: "virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) = None \ virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) = Some (pp (virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm))))" by (metis (no_types) option.exhaust) have f5: "virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) = virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))" using a1 by (meson add_data_cache_low_equal virt_to_phys_low_equal) { assume "Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s1 addr data bm)) \ (case Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s1 addr data bm)) else None)" then have "None = (if mmu_writable (get_acc_flag (snd (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s2 addr data bm)) else None)" by fastforce moreover { assume "(if mmu_writable (get_acc_flag (snd (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s2 addr data bm)) else None) \ (case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None)" then have "(case Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) of None \ if mmu_writable (get_acc_flag (snd (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s1 addr data bm)) else None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None) \ (case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None)" using f3 by simp then have "Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) \ virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) \ (if mmu_writable (get_acc_flag (snd (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s1 addr data bm)) else None) \ None" proof - have "(case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ if mmu_writable (get_acc_flag (snd (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s1 addr data bm)) else None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None) = (case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None) \ Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) \ virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) \ (if mmu_writable (get_acc_flag (snd (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s1 addr data bm)) else None) \ None" by simp then show ?thesis using \(case Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) of None \ if mmu_writable (get_acc_flag (snd (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s1 addr data bm)) else None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None) \ (case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None)\ by force qed moreover { assume "Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) \ virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))" then have "virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) \ Some (pp (virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm))))" using f5 by simp } ultimately have "virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) \ Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) \ virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) \ Some (pp (virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm))))" using a2 by force } ultimately have "virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) = Some (pp (virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)))) \ virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) = Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) \ (case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None) = None" by fastforce } then have "virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) = Some (pp (virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)))) \ virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) = Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) \ (case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None) = None" using a2 by force then show "(case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None) = None" using f5 f4 by force qed lemma memory_write_asi_low_equal_none: "low_equal s1 s2 \ memory_write_asi 10 addr bm data s1 = None \ memory_write_asi 10 addr bm data s2 = None" apply (simp add: memory_write_asi_def) by (simp add: store_word_mem_low_equal_none) lemma memory_write_low_equal_none: "low_equal s1 s2 \ memory_write 10 addr bm data s1 = None \ memory_write 10 addr bm data s2 = None" apply (simp add: memory_write_def) by (metis map_option_case memory_write_asi_low_equal_none option.map_disc_iff) lemma memory_write_low_equal_none2: "low_equal s1 s2 \ memory_write 10 addr bm data s2 = None \ memory_write 10 addr bm data s1 = None" apply (simp add: memory_write_def) by (metis low_equal_com memory_write_def memory_write_low_equal_none) lemma mem_context_val_9_unchanged: "mem_context_val 9 addr1 (mem s1) = mem_context_val 9 addr1 ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None)))" apply (simp add: mem_context_val_def) by (simp add: Let_def) lemma mem_context_val_w32_9_unchanged: "mem_context_val_w32 9 addr1 (mem s1) = mem_context_val_w32 9 addr1 ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None)))" apply (simp add: mem_context_val_w32_def) apply (simp add: Let_def) by (metis mem_context_val_9_unchanged) lemma ptd_lookup_unchanged_4: "ptd_lookup va ptp (mem s1) 4 = ptd_lookup va ptp ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 4" by auto lemma ptd_lookup_unchanged_3: "ptd_lookup va ptp (mem s1) 3 = ptd_lookup va ptp ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 3" proof (cases "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) (mem s1) = None") case True then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) (mem s1) = None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = None" using mem_context_val_w32_9_unchanged by metis then show ?thesis by auto next case False then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) (mem s1) \ None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None" using mem_context_val_w32_9_unchanged by metis then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) (mem s1) \ None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None \ (\y. (mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) (mem s1) = Some y) \ (mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None)))= Some y))" using mem_context_val_w32_9_unchanged by metis then show ?thesis apply auto by (simp add: Let_def) qed lemma ptd_lookup_unchanged_2: "ptd_lookup va ptp (mem s1) 2 = ptd_lookup va ptp ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 2" proof (cases "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) (mem s1) = None") case True then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) (mem s1) = None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = None" using mem_context_val_w32_9_unchanged by metis then show ?thesis by auto next case False then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) (mem s1) \ None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None" using mem_context_val_w32_9_unchanged by metis then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) (mem s1) \ None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None \ (\y. (mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) (mem s1) = Some y) \ (mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None)))= Some y))" using mem_context_val_w32_9_unchanged by metis then show ?thesis apply auto using ptd_lookup_unchanged_3 unfolding Let_def by auto qed lemma ptd_lookup_unchanged_1: "ptd_lookup va ptp (mem s1) 1 = ptd_lookup va ptp ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1" proof (cases "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) (mem s1) = None") case True then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) (mem s1) = None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = None" using mem_context_val_w32_9_unchanged by metis then show ?thesis by auto next case False then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) (mem s1) \ None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None" using mem_context_val_w32_9_unchanged by metis then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) (mem s1) \ None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None \ (\y. (mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) (mem s1) = Some y) \ (mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None)))= Some y))" using mem_context_val_w32_9_unchanged by metis then show ?thesis apply auto using ptd_lookup_unchanged_2 unfolding Let_def proof - fix y :: word32 have "(y AND 3 \ 0 \ y AND 3 = 0 \ (y AND 3 \ 1 \ ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) = None) \ (y AND 3 = 1 \ y AND 3 \ 2 \ None = Some ((ucast (ucast (y >> 8)::word24) << 12) OR (ucast (ucast va::word12)::word36), ucast y::word8))) \ (y AND 3 = 0 \ (y AND 3 \ 1 \ (y AND 3 \ 0 \ ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) = None) \ (y AND 3 = 0 \ (y AND 3 \ 1 \ ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) = ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1)) \ (y AND 3 = 1 \ (y AND 3 \ 2 \ ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) = Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y)) \ (y AND 3 = 2 \ ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) = None)))) \ (y AND 3 = 1 \ (y AND 3 \ 2 \ (y AND 3 \ 0 \ None = Some ((ucast (ucast (y >> 8)::word24) << 12) OR (ucast (ucast va::word12)::word36), ucast y::word8)) \ (y AND 3 = 0 \ (y AND 3 \ 1 \ ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) = Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y)) \ (y AND 3 = 1 \ y AND 3 = 2 \ None = Some ((ucast (ucast (y >> 8)::word24) << 12) OR (ucast (ucast va::word12)::word36), ucast y::word8)))) \ (y AND 3 = 2 \ y AND 3 = 0 \ (y AND 3 \ 1 \ ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) = None) \ (y AND 3 = 1 \ y AND 3 \ 2 \ None = Some ((ucast (ucast (y >> 8)::word24) << 12) OR (ucast (ucast va::word12)::word36), ucast y::word8))))) \ (\w. mem s1 w = ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) w)" by (metis (no_types) One_nat_def Suc_1 Suc_eq_plus1 ptd_lookup_unchanged_2) then show "(if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) = (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None)" proof - have f1: "2 = Suc 0 + 1" by (metis One_nat_def Suc_1 Suc_eq_plus1) { assume "y AND 3 = 1" moreover { assume "y AND 3 = 1 \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None)" have "y AND 3 = 1 \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) \ ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) = (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None)" by presburger moreover { assume "y AND 3 = 1 \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) \ ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1)" then have "y AND 3 = 1 \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) \ ptd_lookup va (y AND 4294967292) (mem s1) 2" by (metis One_nat_def Suc_1 Suc_eq_plus1 ptd_lookup_unchanged_2) then have ?thesis using f1 by auto } ultimately have ?thesis by blast } ultimately have ?thesis by blast } then show ?thesis by presburger qed qed qed lemma virt_to_phys_unchanged_sub1: assumes a1: "(let context_table_entry = (v1 >> 11 << 11) OR (v2 AND 511 << 2) in Let (mem_context_val_w32 (word_of_int 9) (ucast context_table_entry) (mem s1)) (case_option None (\lvl1_page_table. ptd_lookup va lvl1_page_table (mem s1) 1))) = (let context_table_entry = (v1 >> 11 << 11) OR (v2 AND 511 << 2) in Let (mem_context_val_w32 (word_of_int 9) (ucast context_table_entry) (mem s2)) (case_option None (\lvl1_page_table. ptd_lookup va lvl1_page_table (mem s2) 1)))" shows "(let context_table_entry = (v1 >> 11 << 11) OR (v2 AND 511 << 2) in Let (mem_context_val_w32 (word_of_int 9) (ucast context_table_entry) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None)))) (case_option None (\lvl1_page_table. ptd_lookup va lvl1_page_table ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1))) = (let context_table_entry = (v1 >> 11 << 11) OR (v2 AND 511 << 2) in Let (mem_context_val_w32 (word_of_int 9) (ucast context_table_entry) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None)))) (case_option None (\lvl1_page_table. ptd_lookup va lvl1_page_table ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) 1)))" proof - from a1 have "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) (mem s1) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s1) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) (mem s2) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s2) 1)" unfolding Let_def by auto then have "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s1) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s2) 1)" using mem_context_val_w32_9_unchanged by (metis word_numeral_alt) then have "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) 1)" using ptd_lookup_unchanged_1 proof - obtain ww :: "word32 option \ word32" where f1: "\z. (z = None \ z = Some (ww z)) \ (z \ None \ (\w. z \ Some w))" by moura then have f2: "(mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = None \ mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = Some (ww (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None)))))) \ (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None \ (\w. mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ Some w))" by blast then have f3: "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1) \ None \ (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w (mem s1) 1)" by (metis (no_types) \\val va s1 ptp addr. ptd_lookup va ptp (mem s1) 1 = ptd_lookup va ptp ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1\ option.case(2) option.simps(4)) have f4: "mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = Some (ww (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))))) \ mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = Some (ww (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))))) \ (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1) = (case Some (ww (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))))) of None \ None | Some w \ ptd_lookup va w ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) 1)" by (metis (no_types) \(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s1) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s2) 1)\ \\val va s1 ptp addr. ptd_lookup va ptp (mem s1) 1 = ptd_lookup va ptp ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1\ option.case(2)) have f5: "(mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = None \ mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = Some (ww (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None)))))) \ (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) \ None \ (\w. mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) \ Some w))" using f1 by blast { assume "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1) \ (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) 1)" { assume "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w (mem s2) 1) \ None \ (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w (mem s2) 1) \ None" then have "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w (mem s2) 1) \ None \ mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None" by (metis (no_types) \(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s1) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s2) 1)\ option.simps(4)) then have ?thesis using f5 f4 f2 by force } then have ?thesis using f5 f3 by (metis (no_types) \(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s1) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s2) 1)\ \\val va s1 ptp addr. ptd_lookup va ptp (mem s1) 1 = ptd_lookup va ptp ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1\ option.case(2) option.simps(4)) } then show ?thesis by blast qed then show ?thesis unfolding Let_def by auto qed lemma virt_to_phys_unchanged: assumes a1: "(\va. virt_to_phys va (mmu s2) (mem s1) = virt_to_phys va (mmu s2) (mem s2))" shows "(\va. virt_to_phys va (mmu s2) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = virt_to_phys va (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))))" proof (cases "registers (mmu s2) CR AND 1 \ 0") case True then have f1: "registers (mmu s2) CR AND 1 \ 0" by auto then show ?thesis proof (cases "mmu_reg_val (mmu s2) 256 = None") case True then show ?thesis by (simp add: virt_to_phys_def) next case False then have f2: "mmu_reg_val (mmu s2) 256 \ None" by auto then show ?thesis proof (cases "mmu_reg_val (mmu s2) 512 = None") case True then show ?thesis using f1 f2 apply (simp add: virt_to_phys_def) by auto next case False then show ?thesis using f1 f2 a1 apply (simp add: virt_to_phys_def) apply clarify using virt_to_phys_unchanged_sub1 by fastforce qed qed next case False then show ?thesis by (simp add: virt_to_phys_def) qed lemma virt_to_phys_unchanged2_sub1: "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) (mem s2) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s2) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) 1)" proof (cases "mem_context_val_w32 9 (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) (mem s2) = None") case True then have "mem_context_val_w32 9 (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) (mem s2) = None \ mem_context_val_w32 9 (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = None" using mem_context_val_w32_9_unchanged by metis then show ?thesis by auto next case False then have "mem_context_val_w32 9 (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) (mem s2) \ None \ (\y. mem_context_val_w32 9 (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) (mem s2) = Some y \ mem_context_val_w32 9 (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = Some y)" using mem_context_val_w32_9_unchanged by metis then show ?thesis using ptd_lookup_unchanged_1 by fastforce qed lemma virt_to_phys_unchanged2: "virt_to_phys va (mmu s2) (mem s2) = virt_to_phys va (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None)))" proof (cases "registers (mmu s2) CR AND 1 \ 0") case True then have f1: "registers (mmu s2) CR AND 1 \ 0" by auto then show ?thesis proof (cases "mmu_reg_val (mmu s2) 256 = None") case True then show ?thesis by (simp add: virt_to_phys_def) next case False then have f2: "mmu_reg_val (mmu s2) 256 \ None" by auto then show ?thesis proof (cases "mmu_reg_val (mmu s2) 512 = None") case True then show ?thesis using f1 f2 apply (simp add: virt_to_phys_def) by auto next case False then show ?thesis using f1 f2 apply (simp add: virt_to_phys_def) apply clarify unfolding Let_def using virt_to_phys_unchanged2_sub1 by auto qed qed next case False then show ?thesis by (simp add: virt_to_phys_def) qed lemma virt_to_phys_unchanged_low_equal: assumes a1: "low_equal s1 s2" shows "(\va. virt_to_phys va (mmu s2) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = virt_to_phys va (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))))" using a1 apply (simp add: low_equal_def) using virt_to_phys_unchanged by metis lemma mmu_low_equal: "low_equal s1 s2 \ mmu s1 = mmu s2" by (simp add: low_equal_def) lemma mem_val_alt_8_unchanged0: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 8 (pa AND 68719476732) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 8 (pa AND 68719476732) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_unchanged1: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 8 ((pa AND 68719476732) + 1) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 8 ((pa AND 68719476732) + 1) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_unchanged2: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 8 ((pa AND 68719476732) + 2) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 8 ((pa AND 68719476732) + 2) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_unchanged3: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 8 ((pa AND 68719476732) + 3) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 8 ((pa AND 68719476732) + 3) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_unchanged: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 8 (pa AND 68719476732) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 8 (pa AND 68719476732) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) \ mem_val_alt 8 ((pa AND 68719476732) + 1) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 8 ((pa AND 68719476732) + 1) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) \ mem_val_alt 8 ((pa AND 68719476732) + 2) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 8 ((pa AND 68719476732) + 2) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) \ mem_val_alt 8 ((pa AND 68719476732) + 3) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 8 ((pa AND 68719476732) + 3) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" using a1 mem_val_alt_8_unchanged0 mem_val_alt_8_unchanged1 mem_val_alt_8_unchanged2 mem_val_alt_8_unchanged3 by blast lemma mem_val_w32_8_unchanged: assumes a1: "mem_equal s1 s2 a" shows "mem_val_w32 8 a (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_w32 8 a (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_w32_def) apply (simp add: Let_def) using mem_val_alt_8_unchanged a1 apply auto apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce by fastforce lemma load_word_mem_8_unchanged: assumes a1: "low_equal s1 s2 \ load_word_mem s1 addra 8 = load_word_mem s2 addra 8" shows "load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 8 = load_word_mem (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) addra 8" proof (cases "virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = None") case True then have "virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = None \ virt_to_phys addra (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = None" using a1 apply (auto simp add: mmu_low_equal) using a1 virt_to_phys_unchanged_low_equal by metis then show ?thesis by (simp add: load_word_mem_def) next case False then have "\p. virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = Some p \ virt_to_phys addra (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = Some p" using a1 apply (auto simp add: mmu_low_equal) using a1 virt_to_phys_unchanged_low_equal by metis then have "\p. virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = Some p \ virt_to_phys addra (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = Some p \ virt_to_phys addra (mmu s1) (mem s1) = Some p \ virt_to_phys addra (mmu s2) (mem s2) = Some p" using virt_to_phys_unchanged2 by metis then show ?thesis using a1 apply (simp add: load_word_mem_def) apply auto apply (simp add: low_equal_def) apply (simp add: user_accessible_def) using mem_val_w32_8_unchanged a1 user_accessible_8 by (metis snd_conv) qed lemma load_word_mem_select_8: assumes a1: "fst (case load_word_mem s1 addra 8 of None \ (None, s1) | Some w \ (Some w, add_instr_cache s1 addra w 15)) = fst (case load_word_mem s2 addra 8 of None \ (None, s2) | Some w \ (Some w, add_instr_cache s2 addra w 15))" shows "load_word_mem s1 addra 8 = load_word_mem s2 addra 8" using a1 by (metis (mono_tags, lifting) fst_conv not_None_eq option.simps(4) option.simps(5)) lemma memory_read_8_unchanged: assumes a1: "low_equal s1 s2 \ fst (memory_read 8 addra s1) = fst (memory_read 8 addra s2)" shows "fst (memory_read 8 addra (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\)) = fst (memory_read 8 addra (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\))" proof (cases "sys_reg s1 CCR AND 1 = 0") case True then have "sys_reg s1 CCR AND 1 = 0 \ sys_reg s2 CCR AND 1 = 0" using a1 sys_reg_low_equal by fastforce then show ?thesis using a1 apply (simp add: memory_read_def) using load_word_mem_8_unchanged by blast next case False then have f1: "sys_reg s1 CCR AND 1 \ 0 \ sys_reg s2 CCR AND 1 \ 0" using a1 sys_reg_low_equal by fastforce then show ?thesis using a1 proof (cases "load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 8 = None") case True then have "load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 8 = None \ load_word_mem (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) addra 8 = None" using a1 f1 apply (simp add: memory_read_def) apply clarsimp using load_word_mem_select_8 load_word_mem_8_unchanged by fastforce then show ?thesis by (simp add: memory_read_def) next case False then have "\y. load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 8 = Some y" by auto then have "\y. load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 8 = Some y \ load_word_mem (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) addra 8 = Some y" using a1 f1 apply (simp add: memory_read_def) apply clarsimp using load_word_mem_select_8 load_word_mem_8_unchanged by fastforce then show ?thesis using a1 f1 apply (simp add: memory_read_def) by auto qed qed lemma mem_val_alt_mod: assumes a1: "addr1 \ addr2" shows "mem_val_alt 10 addr1 s = mem_val_alt 10 addr1 (s\mem := (mem s)(10 := mem s 10(addr2 \ val), 11 := (mem s 11)(addr2 := None))\)" using a1 apply (simp add: mem_val_alt_def) by (simp add: Let_def) lemma mem_val_alt_mod2: "mem_val_alt 10 addr (s\mem := (mem s)(10 := mem s 10(addr \ val), 11 := (mem s 11)(addr := None))\) = Some val" by (simp add: mem_val_alt_def) lemma mem_val_alt_10_unchanged0: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 10 (pa AND 68719476732) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 10 (pa AND 68719476732) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_unchanged1: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 10 ((pa AND 68719476732) + 1) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 10 ((pa AND 68719476732) + 1) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_unchanged2: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 10 ((pa AND 68719476732) + 2) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 10 ((pa AND 68719476732) + 2) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_unchanged3: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 10 ((pa AND 68719476732) + 3) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 10 ((pa AND 68719476732) + 3) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_unchanged: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 10 (pa AND 68719476732) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 10 (pa AND 68719476732) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) \ mem_val_alt 10 ((pa AND 68719476732) + 1) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 10 ((pa AND 68719476732) + 1) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) \ mem_val_alt 10 ((pa AND 68719476732) + 2) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 10 ((pa AND 68719476732) + 2) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) \ mem_val_alt 10 ((pa AND 68719476732) + 3) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 10 ((pa AND 68719476732) + 3) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" using a1 mem_val_alt_10_unchanged0 mem_val_alt_10_unchanged1 mem_val_alt_10_unchanged2 mem_val_alt_10_unchanged3 by blast lemma mem_val_w32_10_unchanged: assumes a1: "mem_equal s1 s2 a" shows "mem_val_w32 10 a (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_w32 10 a (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_w32_def) apply (simp add: Let_def) using mem_val_alt_10_unchanged a1 apply auto apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce by fastforce lemma is_accessible: "low_equal s1 s2 \ virt_to_phys addra (mmu s1) (mem s1) = Some (a, b) \ virt_to_phys addra (mmu s2) (mem s2) = Some (a, b) \ mmu_readable (get_acc_flag b) 10 \ mem_equal s1 s2 a" apply (simp add: low_equal_def) apply (simp add: user_accessible_def) by fastforce lemma load_word_mem_10_unchanged: assumes a1: "low_equal s1 s2 \ load_word_mem s1 addra 10 = load_word_mem s2 addra 10" shows "load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 10 = load_word_mem (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) addra 10" proof (cases "virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = None") case True then have "virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = None \ virt_to_phys addra (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = None" using a1 apply (auto simp add: mmu_low_equal) using a1 virt_to_phys_unchanged_low_equal by metis then show ?thesis by (simp add: load_word_mem_def) next case False then have "\p. virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = Some p \ virt_to_phys addra (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = Some p" using a1 apply (auto simp add: mmu_low_equal) using a1 virt_to_phys_unchanged_low_equal by metis then have "\p. virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = Some p \ virt_to_phys addra (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = Some p \ virt_to_phys addra (mmu s1) (mem s1) = Some p \ virt_to_phys addra (mmu s2) (mem s2) = Some p" using virt_to_phys_unchanged2 by metis then show ?thesis using a1 apply (simp add: load_word_mem_def) apply auto apply (simp add: low_equal_def) apply (simp add: user_accessible_def) using mem_val_w32_10_unchanged a1 by metis qed lemma load_word_mem_select_10: assumes a1: "fst (case load_word_mem s1 addra 10 of None \ (None, s1) | Some w \ (Some w, add_data_cache s1 addra w 15)) = fst (case load_word_mem s2 addra 10 of None \ (None, s2) | Some w \ (Some w, add_data_cache s2 addra w 15))" shows "load_word_mem s1 addra 10 = load_word_mem s2 addra 10" using a1 by (metis (mono_tags, lifting) fst_conv not_None_eq option.simps(4) option.simps(5)) lemma memory_read_10_unchanged: assumes a1: "low_equal s1 s2 \ fst (memory_read 10 addra s1) = fst (memory_read 10 addra s2)" shows "fst (memory_read 10 addra (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\)) = fst (memory_read 10 addra (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\))" proof (cases "sys_reg s1 CCR AND 1 = 0") case True then have "sys_reg s1 CCR AND 1 = 0 \ sys_reg s2 CCR AND 1 = 0" using a1 sys_reg_low_equal by fastforce then show ?thesis using a1 apply (simp add: memory_read_def) using load_word_mem_10_unchanged by blast next case False then have f1: "sys_reg s1 CCR AND 1 \ 0 \ sys_reg s2 CCR AND 1 \ 0" using a1 sys_reg_low_equal by fastforce then show ?thesis using a1 proof (cases "load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 10 = None") case True then have "load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 10 = None \ load_word_mem (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) addra 10 = None" using a1 f1 apply (simp add: memory_read_def) apply clarsimp using load_word_mem_select_10 load_word_mem_10_unchanged by fastforce then show ?thesis by (simp add: memory_read_def) next case False then have "\y. load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 10 = Some y" by auto then have "\y. load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 10 = Some y \ load_word_mem (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) addra 10 = Some y" using a1 f1 apply (simp add: memory_read_def) apply clarsimp using load_word_mem_select_10 load_word_mem_10_unchanged by fastforce then show ?thesis using a1 f1 apply (simp add: memory_read_def) by auto qed qed lemma state_mem_mod_1011_low_equal_sub1: assumes a1: "(\va. virt_to_phys va (mmu s2) (mem s1) = virt_to_phys va (mmu s2) (mem s2)) \ (\pa. (\va b. virt_to_phys va (mmu s2) (mem s2) = Some (pa, b) \ mmu_readable (get_acc_flag b) 10) \ mem_equal s1 s2 pa) \ mmu s1 = mmu s2 \ virt_to_phys va (mmu s2) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = Some (pa, b) \ mmu_readable (get_acc_flag b) 10" shows "mem_equal s1 s2 pa" proof - have "virt_to_phys va (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = Some (pa, b)" using a1 by auto then have "virt_to_phys va (mmu s1) (mem s1) = Some (pa, b)" using virt_to_phys_unchanged2 by metis then have "virt_to_phys va (mmu s2) (mem s2) = Some (pa, b)" using a1 by auto then show ?thesis using a1 by auto qed lemma mem_equal_unchanged: assumes a1: "mem_equal s1 s2 pa" shows "mem_equal (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) pa" using a1 apply (simp add: mem_equal_def) by auto lemma state_mem_mod_1011_low_equal: assumes a1: "low_equal s1 s2 \ t1 = s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\ \ t2 = s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\" shows "low_equal t1 t2" using a1 apply (simp add: low_equal_def) apply (simp add: user_accessible_def) apply auto apply (simp add: assms virt_to_phys_unchanged_low_equal) using state_mem_mod_1011_low_equal_sub1 mem_equal_unchanged apply metis apply (metis virt_to_phys_unchanged2) using state_mem_mod_1011_low_equal_sub1 mem_equal_unchanged by metis lemma mem_mod_low_equal: assumes a1: "low_equal s1 s2 \ t1 = (mem_mod 10 addr val s1) \ t2 = (mem_mod 10 addr val s2)" shows "low_equal t1 t2" using a1 apply (simp add: mem_mod_def) by (auto intro: state_mem_mod_1011_low_equal) lemma mem_mod_w32_low_equal: assumes a1: "low_equal s1 s2 \ t1 = mem_mod_w32 10 a bm data s1 \ t2 = mem_mod_w32 10 a bm data s2" shows "low_equal t1 t2" using a1 apply (simp add: mem_mod_w32_def) apply (simp add: Let_def) by (meson mem_mod_low_equal) lemma store_word_mem_low_equal: assumes a1: "low_equal s1 s2 \ Some t1 = store_word_mem s1 addr data bm 10 \ Some t2 = store_word_mem s2 addr data bm 10" shows "low_equal t1 t2" using a1 apply (simp add: store_word_mem_def) apply (auto simp add: virt_to_phys_low_equal) apply (case_tac "virt_to_phys addr (mmu s2) (mem s2) = None") apply auto apply (case_tac "mmu_writable (get_acc_flag b) 10") apply auto using mem_mod_w32_low_equal by blast lemma memory_write_asi_low_equal: assumes a1: "low_equal s1 s2 \ Some t1 = memory_write_asi 10 addr bm data s1 \ Some t2 = memory_write_asi 10 addr bm data s2" shows "low_equal t1 t2" using a1 apply (simp add: memory_write_asi_def) by (meson add_data_cache_low_equal store_word_mem_low_equal) lemma store_barrier_pending_mod_low_equal: assumes a1: "low_equal s1 s2 \ t1 = store_barrier_pending_mod False s1 \ t2 = store_barrier_pending_mod False s2" shows "low_equal t1 t2" using a1 apply (simp add: store_barrier_pending_mod_def) apply clarsimp using a1 apply (auto simp add: state_var_low_equal) by (auto intro: state_var2_low_equal) lemma memory_write_low_equal: assumes a1: "low_equal s1 s2 \ Some t1 = memory_write 10 addr bm data s1 \ Some t2 = memory_write 10 addr bm data s2" shows "low_equal t1 t2" apply (case_tac "memory_write_asi 10 addr bm data s1 = None") using a1 apply (simp add: memory_write_def) apply (case_tac "memory_write_asi 10 addr bm data s2 = None") apply (meson assms low_equal_com memory_write_asi_low_equal_none) using a1 apply (simp add: memory_write_def) apply auto by (metis memory_write_asi_low_equal store_barrier_pending_mod_low_equal) lemma memory_write_low_equal2: assumes a1: "low_equal s1 s2 \ Some t1 = memory_write 10 addr bm data s1" shows "\t2. Some t2 = memory_write 10 addr bm data s2" using a1 apply (simp add: memory_write_def) apply auto by (metis (full_types) memory_write_def memory_write_low_equal_none2 not_None_eq) lemma store_sub2_low_equal_sub1: assumes a1: "low_equal s1 s2 \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s1 = Some y \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s2 = Some ya" shows "low_equal (y\traps := insert data_access_exception (traps y)\) (ya\traps := insert data_access_exception (traps ya)\)" proof - from a1 have f1: "low_equal y ya" using memory_write_low_equal by metis then have "traps y = traps ya" by (simp add: low_equal_def) then show ?thesis using f1 mod_trap_low_equal by fastforce qed lemma store_sub2_low_equal_sub2: assumes a1: "low_equal s1 s2 \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s1 = Some y \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s2 = Some ya \ memory_write 10 (addr + 4) 15 (user_reg_val curr_win (rd OR 1) y) y = None \ memory_write 10 (addr + 4) 15 (user_reg_val curr_win (rd OR 1) ya) ya = Some yb" shows "False" proof - from a1 have f1: "low_equal y ya" using memory_write_low_equal by metis then have "(user_reg_val curr_win (rd OR 1) y) = (user_reg_val curr_win (rd OR 1) ya)" by (simp add: low_equal_def user_reg_val_def) then show ?thesis using a1 using f1 memory_write_low_equal_none by fastforce qed lemma store_sub2_low_equal_sub3: assumes a1: "low_equal s1 s2 \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s1 = Some y \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s2 = Some ya \ memory_write 10 (addr + 4) 15 (user_reg_val curr_win (rd OR 1) y) y = Some yb \ memory_write 10 (addr + 4) 15 (user_reg_val curr_win (rd OR 1) ya) ya = None" shows "False" proof - from a1 have f1: "low_equal y ya" using memory_write_low_equal by metis then have "(user_reg_val curr_win (rd OR 1) y) = (user_reg_val curr_win (rd OR 1) ya)" by (simp add: low_equal_def user_reg_val_def) then show ?thesis using a1 using f1 memory_write_low_equal_none2 by fastforce qed lemma store_sub2_low_equal_sub4: assumes a1: "low_equal s1 s2 \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s1 = Some y \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s2 = Some ya \ memory_write 10 (addr + 4) 15 (user_reg_val curr_win (rd OR 1) y) y = Some yb \ memory_write 10 (addr + 4) 15 (user_reg_val curr_win (rd OR 1) ya) ya = Some yc" shows "low_equal yb yc" proof - from a1 have f1: "low_equal y ya" using memory_write_low_equal by metis then have "(user_reg_val curr_win (rd OR 1) y) = (user_reg_val curr_win (rd OR 1) ya)" by (simp add: low_equal_def user_reg_val_def) then show ?thesis using a1 f1 by (metis memory_write_low_equal) qed lemma store_sub2_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (store_sub2 instr curr_win rd 10 addr s1)) \ t2 = snd (fst (store_sub2 instr curr_win rd 10 addr s2))" shows "low_equal t1 t2" proof (cases "memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s1) s1 = None") case True then have "memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s1) s1 = None \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s2 = None" using a1 by (metis memory_write_low_equal_none st_data0_low_equal) then show ?thesis using a1 apply (simp add: store_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold return_def) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) using mod_trap_low_equal traps_low_equal by fastforce next case False then have f1: "memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s1) s1 \ None \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s2 \ None" using a1 by (metis memory_write_low_equal_none2 st_data0_low_equal) then show ?thesis proof (cases "(fst instr) \ {load_store_type STD,load_store_type STDA}") case True then show ?thesis using a1 f1 apply (simp add: store_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) apply (simp add: return_def) apply (simp add: bind_def case_prod_unfold) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: case_prod_unfold bind_def h1_def h2_def Let_def simpler_modify_def) apply (simp add: simpler_gets_def) apply auto apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply (simp add: st_data0_low_equal) apply (simp add: store_sub2_low_equal_sub1) apply (simp add: st_data0_low_equal) using store_sub2_low_equal_sub2 apply blast apply (simp add: st_data0_low_equal) using store_sub2_low_equal_sub3 apply blast apply (simp add: st_data0_low_equal) using store_sub2_low_equal_sub4 apply blast apply (simp add: st_data0_low_equal) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) using store_sub2_low_equal_sub1 apply blast apply (simp add: st_data0_low_equal) using store_sub2_low_equal_sub2 apply blast apply (simp add: st_data0_low_equal) using store_sub2_low_equal_sub3 apply blast apply (simp add: st_data0_low_equal) using store_sub2_low_equal_sub4 by blast next case False then show ?thesis using a1 f1 apply (simp add: store_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) apply (simp add: return_def) apply (simp add: bind_def case_prod_unfold) apply clarsimp apply (simp add: simpler_modify_def) apply (simp add: st_data0_low_equal) using memory_write_low_equal by metis qed qed lemma store_sub1_low_equal: assumes a1: "low_equal s1 s2 \ (fst instr = load_store_type STB \ fst instr = load_store_type STH \ fst instr = load_store_type ST \ fst instr = load_store_type STD) \ t1 = snd (fst (store_sub1 instr rd 0 s1)) \ t2 = snd (fst (store_sub1 instr rd 0 s2))" shows "low_equal t1 t2" proof (cases "(fst instr = load_store_type STH \ fst instr = load_store_type STHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word1) \ 0") case True then have "((fst instr = load_store_type STH \ fst instr = load_store_type STHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word1) \ 0) \ ((fst instr = load_store_type STH \ fst instr = load_store_type STHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word1) \ 0)" by (metis (mono_tags, lifting) assms get_addr_low_equal) then show ?thesis using a1 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: get_curr_win3_low_equal) by (auto intro: get_curr_win2_low_equal mod_trap_low_equal) next case False then have f1: "\ ((fst instr = load_store_type STH \ fst instr = load_store_type STHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word1) \ 0) \ \ ((fst instr = load_store_type STH \ fst instr = load_store_type STHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word1) \ 0)" by (metis (mono_tags, lifting) assms get_addr_low_equal) then show ?thesis proof (cases "(fst instr \ {load_store_type ST,load_store_type STA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) \ 0") case True then have "(fst instr \ {load_store_type ST,load_store_type STA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) \ 0 \ (fst instr \ {load_store_type ST,load_store_type STA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word2) \ 0" by (metis (mono_tags, lifting) assms get_addr_low_equal) then show ?thesis using a1 f1 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: get_curr_win3_low_equal) by (auto intro: get_curr_win2_low_equal mod_trap_low_equal) next case False then have "\((fst instr \ {load_store_type ST,load_store_type STA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) \ 0) \ \((fst instr \ {load_store_type ST,load_store_type STA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word2) \ 0)" by (metis (mono_tags, lifting) assms get_addr_low_equal) then have f2: "\((fst instr = load_store_type ST \ fst instr = load_store_type STA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) \ 0) \ \((fst instr = load_store_type ST \ fst instr = load_store_type STA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word2) \ 0)" by auto then show ?thesis proof (cases "(fst instr \ {load_store_type STD,load_store_type STDA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) \ 0") case True then have "(fst instr \ {load_store_type STD,load_store_type STDA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) \ 0 \ (fst instr \ {load_store_type STD,load_store_type STDA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word3) \ 0" by (metis (mono_tags, lifting) assms get_addr_low_equal) then show ?thesis using a1 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply auto apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply (simp add: get_curr_win3_low_equal) by (auto intro: get_curr_win2_low_equal mod_trap_low_equal) next case False then have "\ (fst instr \ {load_store_type STD, load_store_type STDA} \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) \ 0) \ \ (fst instr \ {load_store_type STD, load_store_type STDA} \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word3) \ 0)" by (metis (mono_tags, lifting) assms get_addr_low_equal) then have f3: "\ ((fst instr = load_store_type STD \ fst instr = load_store_type STDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) \ 0) \ \ ((fst instr = load_store_type STD \ fst instr = load_store_type STDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word3) \ 0)" by auto show ?thesis using a1 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (unfold case_prod_beta) apply (simp add: f1 f2 f3) apply (simp_all add: st_asi_def) using a1 apply clarsimp apply (simp add: get_curr_win_low_equal get_addr2_low_equal) by (metis store_sub2_low_equal get_curr_win2_low_equal) qed qed qed lemma store_instr_low_equal: assumes a1: "low_equal s1 s2 \ (fst instr = load_store_type STB \ fst instr = load_store_type STH \ fst instr = load_store_type ST \ fst instr = load_store_type STA \ fst instr = load_store_type STD) \ ((ucast (get_S (cpu_reg_val PSR s1)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s2)))::word1) = 0 \ t1 = snd (fst (store_instr instr s1)) \ t2 = snd (fst (store_instr instr s2))" shows "low_equal t1 t2" proof - have "get_S (cpu_reg_val PSR s1) = 0 \ get_S (cpu_reg_val PSR s2) = 0" using a1 by (simp add: ucast_id) then show ?thesis using a1 apply (simp add: store_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) apply clarsimp apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply (simp add: traps_low_equal) by (auto intro: mod_trap_low_equal store_sub1_low_equal) qed lemma sethi_low_equal: "low_equal s1 s2 \ t1 = snd (fst (sethi_instr instr s1)) \ t2 = snd (fst (sethi_instr instr s2)) \ low_equal t1 t2" apply (simp add: sethi_instr_def) apply (simp add: Let_def) apply (case_tac "get_operand_w5 (snd instr ! Suc 0) \ 0") apply auto apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: get_curr_win_low_equal) using get_curr_win2_low_equal write_reg_low_equal apply metis by (simp add: return_def) lemma nop_low_equal: "low_equal s1 s2 \ t1 = snd (fst (nop_instr instr s1)) \ t2 = snd (fst (nop_instr instr s2)) \ low_equal t1 t2" apply (simp add: nop_instr_def) by (simp add: return_def) lemma logical_instr_sub1_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (logical_instr_sub1 instr_name result s1)) \ t2 = snd (fst (logical_instr_sub1 instr_name result s2))" shows "low_equal t1 t2" proof (cases "instr_name = logic_type ANDcc \ instr_name = logic_type ANDNcc \ instr_name = logic_type ORcc \ instr_name = logic_type ORNcc \ instr_name = logic_type XORcc \ instr_name = logic_type XNORcc") case True then show ?thesis using a1 apply (simp add: logical_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: logical_new_psr_val_def) using write_cpu_low_equal cpu_reg_val_low_equal by fastforce next case False then show ?thesis using a1 apply (simp add: logical_instr_sub1_def) by (simp add: return_def) qed lemma logical_instr_low_equal: "low_equal s1 s2 \ t1 = snd (fst (logical_instr instr s1)) \ t2 = snd (fst (logical_instr instr s2)) \ low_equal t1 t2" apply (simp add: logical_instr_def) apply (simp add: Let_def simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply auto apply (simp_all add: get_curr_win_low_equal) apply (simp_all add: get_operand2_low_equal) using logical_instr_sub1_low_equal get_operand2_low_equal get_curr_win2_low_equal write_reg_low_equal user_reg_val_low_equal proof - assume a1: "low_equal s1 s2" assume "t2 = snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))))" assume "t1 = snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))" have "\w wa. user_reg_val w wa (snd (fst (get_curr_win () s2))) = user_reg_val w wa (snd (fst (get_curr_win () s1)))" using a1 by (metis (no_types) get_curr_win2_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))) (snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))))" using a1 by (metis (no_types) get_curr_win2_low_equal logical_instr_sub1_low_equal write_reg_low_equal) next assume a2: "low_equal s1 s2" assume "t1 = snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))" assume "t2 = snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))))" have "\w wa. user_reg_val w wa (snd (fst (get_curr_win () s2))) = user_reg_val w wa (snd (fst (get_curr_win () s1)))" using a2 by (metis (no_types) get_curr_win2_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))" proof - have "low_equal (snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))" by (meson a2 get_curr_win2_low_equal logical_instr_sub1_low_equal write_reg_low_equal) then show ?thesis using \\wa w. user_reg_val w wa (snd (fst (get_curr_win () s2))) = user_reg_val w wa (snd (fst (get_curr_win () s1)))\ by presburger qed qed lemma shift_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (shift_instr instr s1)) \ t2 = snd (fst (shift_instr instr s2))" shows "low_equal t1 t2" proof (cases "(fst instr = shift_type SLL) \ (get_operand_w5 ((snd instr)!3) \ 0)") case True then show ?thesis using a1 apply (simp add: shift_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def) apply (simp add: bind_def h1_def h2_def Let_def case_prod_unfold) apply auto apply (simp_all add: get_curr_win_low_equal) proof - assume a1: "low_equal s1 s2" assume "t2 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) << unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s2))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))" assume "t1 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) << unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s1))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))" have "\w wa wb. low_equal (snd (fst (write_reg w wa wb s1))) (snd (fst (write_reg w wa wb s2)))" using a1 by (metis write_reg_low_equal) then show "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) << unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s1))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) << unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s2))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using a1 by (simp add: get_curr_win_def simpler_gets_def user_reg_val_low_equal) next assume a2: "low_equal s1 s2" assume "t1 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) << unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))" assume "t2 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) << unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))" have "\w wa wb. low_equal (snd (fst (write_reg w wa wb s1))) (snd (fst (write_reg w wa wb s2)))" using a2 by (metis write_reg_low_equal) then show "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) << unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) << unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" proof - assume a1: "\w wa wb. low_equal (snd (fst (write_reg w wa wb s1))) (snd (fst (write_reg w wa wb s2)))" have "\u s. fst (get_curr_win u s) = (ucast (get_CWP (cpu_reg_val PSR s))::'a word, s)" by (simp add: get_curr_win_def simpler_gets_def) then show ?thesis using a1 assms user_reg_val_low_equal by fastforce qed qed next case False then have f1: "\((fst instr = shift_type SLL) \ (get_operand_w5 ((snd instr)!3) \ 0))" by auto then show ?thesis proof (cases "(fst instr = shift_type SRL) \ (get_operand_w5 ((snd instr)!3) \ 0)") case True then show ?thesis using a1 f1 apply (simp add: shift_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def) apply (simp add: bind_def h1_def h2_def Let_def case_prod_unfold) apply auto apply (simp_all add: get_curr_win_low_equal) proof - assume a1: "low_equal s1 s2" assume "t2 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s2))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))" assume "t1 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s1))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))" have "\u s. fst (get_curr_win u s) = (ucast (get_CWP (cpu_reg_val PSR s))::'a word, s)" by (simp add: get_curr_win_def simpler_gets_def) then show "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s1))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s2))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using a1 user_reg_val_low_equal write_reg_low_equal by fastforce next assume a2: "low_equal s1 s2" assume "t1 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))" assume "t2 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))" have "\u s. fst (get_curr_win u s) = (ucast (get_CWP (cpu_reg_val PSR s))::'a word, s)" by (simp add: get_curr_win_def simpler_gets_def) then show "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using a2 user_reg_val_low_equal write_reg_low_equal by fastforce qed next case False then have f2: "\((fst instr = shift_type SRL) \ (get_operand_w5 ((snd instr)!3) \ 0))" by auto then show ?thesis proof (cases "(fst instr = shift_type SRA) \ (get_operand_w5 ((snd instr)!3) \ 0)") case True then show ?thesis using a1 f1 f2 apply (simp add: shift_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def) apply (simp add: bind_def h1_def h2_def Let_def case_prod_unfold) apply auto apply (simp_all add: get_curr_win_low_equal) proof - assume a1: "low_equal s1 s2" assume "t1 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >>> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s1))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))" assume "t2 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >>> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s2))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))" have "\w wa. user_reg_val wa w (snd (fst (get_curr_win () s1))) = user_reg_val wa w (snd (fst (get_curr_win () s2)))" using a1 by (meson get_curr_win2_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >>> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s1))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >>> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s2))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using a1 by (metis (no_types) get_curr_win2_low_equal write_reg_low_equal) next assume a2: "low_equal s1 s2" assume "t1 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >>> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))" assume "t2 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >>> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))" have "\w wa. user_reg_val wa w (snd (fst (get_curr_win () s1))) = user_reg_val wa w (snd (fst (get_curr_win () s2)))" using a2 by (meson get_curr_win2_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >>> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >>> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using a2 get_curr_win2_low_equal write_reg_low_equal by fastforce qed next case False then show ?thesis using a1 f1 f2 apply (simp add: shift_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def) apply (simp add: bind_def h1_def h2_def Let_def case_prod_unfold) apply (simp add: return_def) using get_curr_win2_low_equal by blast qed qed qed lemma add_instr_sub1_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (add_instr_sub1 instr_name result rs1_val operand2 s1)) \ t2 = snd (fst (add_instr_sub1 instr_name result rs1_val operand2 s2))" shows "low_equal t1 t2" proof (cases "instr_name = arith_type ADDcc \ instr_name = arith_type ADDXcc") case True then show ?thesis using a1 apply (simp add: add_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (clarsimp simp add: cpu_reg_val_low_equal) using write_cpu_low_equal by blast next case False then show ?thesis using a1 apply (simp add: add_instr_sub1_def) by (simp add: return_def) qed lemma add_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (add_instr instr s1)) \ t2 = snd (fst (add_instr instr s2))" shows "low_equal t1 t2" proof - have f1: "low_equal s1 s2 \ t1 = snd (fst (add_instr_sub1 (fst instr) (if fst instr = arith_type ADD \ fst instr = arith_type ADDcc then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 else user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) \ 0 then if fst instr = arith_type ADD \ fst instr = arith_type ADDcc then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 else user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1))))) else user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))))) \ t2 = snd (fst (add_instr_sub1 (fst instr) (if fst instr = arith_type ADD \ fst instr = arith_type ADDcc then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 else user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) \ 0 then if fst instr = arith_type ADD \ fst instr = arith_type ADDcc then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 else user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2))))) else user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))))" using a1 apply (simp add: add_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: case_prod_unfold) then show ?thesis proof (cases "get_operand_w5 (snd instr ! 3) \ 0") case True then have f2: "get_operand_w5 (snd instr ! 3) \ 0" by auto then show ?thesis proof (cases "fst instr = arith_type ADD \ fst instr = arith_type ADDcc") case True then show ?thesis using f1 f2 apply clarsimp proof - assume a1: "low_equal s1 s2" assume "t1 = snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))" assume a2: "t2 = snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))))" have f3: "\is. get_operand2 is s1 = get_operand2 is s2" using a1 by (metis get_operand2_low_equal) have f4: "fst (fst (get_curr_win () s1)) = fst (fst (get_curr_win () s2))" using a1 by (meson get_curr_win_low_equal) have "\s. snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) s + get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) s) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) s + get_operand2 (snd instr) s2) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))) = t2 \ \ low_equal s (snd (fst (get_curr_win () s2)))" using a2 user_reg_val_low_equal by fastforce then show "low_equal (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))) (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))))" using f4 f3 a2 a1 by (metis (no_types) add_instr_sub1_low_equal get_curr_win2_low_equal write_reg_low_equal) qed next case False then show ?thesis using f1 f2 apply clarsimp proof - assume a1: "low_equal s1 s2" have f2: "\s sa sb w wa wb sc. (\ low_equal s sa \ sb \ snd (fst (write_reg w (wa::'a word) wb s)) \ sc \ snd (fst (write_reg w wa wb sa))) \ low_equal sb sc" by (meson write_reg_low_equal) have f3: "gets (\s. ucast (get_CWP (cpu_reg_val PSR s))::'a word) = get_curr_win ()" by (simp add: get_curr_win_def) then have "((ucast (get_CWP (cpu_reg_val PSR s1)), s1), False) = (fst (get_curr_win () s1), snd (get_curr_win () s1))" by (metis (no_types) prod.collapse simpler_gets_def) then have "(ucast (get_CWP (cpu_reg_val PSR s1)), s1) = fst (get_curr_win () s1) \ \ snd (get_curr_win () s1)" by blast then have f4: "ucast (get_CWP (cpu_reg_val PSR s1)) = fst (fst (get_curr_win () s1)) \ s1 = snd (fst (get_curr_win () s1))" by (metis (no_types) prod.collapse prod.simps(1)) have "((ucast (get_CWP (cpu_reg_val PSR s2)), s2), False) = (fst (get_curr_win () s2), snd (get_curr_win () s2))" using f3 by (metis (no_types) prod.collapse simpler_gets_def) then have "(ucast (get_CWP (cpu_reg_val PSR s2)), s2) = fst (get_curr_win () s2) \ \ snd (get_curr_win () s2)" by blast then have f5: "ucast (get_CWP (cpu_reg_val PSR s2)) = fst (fst (get_curr_win () s2)) \ s2 = snd (fst (get_curr_win () s2))" by (metis prod.collapse prod.simps(1)) then have f6: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2))) = low_equal s1 s2" using f4 by presburger have f7: "fst (fst (get_curr_win () s1)) = ucast (get_CWP (cpu_reg_val PSR s1))" using f4 by presburger have f8: "cpu_reg_val PSR s1 = cpu_reg_val PSR s2" using a1 by (meson cpu_reg_val_low_equal) have f9: "user_reg_val (ucast (get_CWP (cpu_reg_val PSR s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) = user_reg_val (ucast (get_CWP (cpu_reg_val PSR s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))" using f6 a1 by (meson user_reg_val_low_equal) have f10: "ucast (get_CWP (cpu_reg_val PSR s2)) = fst (fst (get_curr_win () s2))" using f5 by meson have f11: "\s sa is. \ low_equal (s::'a sparc_state) sa \ get_operand2 is s = get_operand2 is sa" using get_operand2_low_equal by blast then have f12: "user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1))))) = user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))" using f9 f8 f5 f4 a1 by auto then have "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using f10 f8 f6 f4 f2 a1 by simp then show "low_equal (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))) (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))))" using f12 f11 f10 f9 f8 f7 a1 add_instr_sub1_low_equal by fastforce qed qed next case False then have f3: "\ get_operand_w5 (snd instr ! 3) \ 0" by auto then show ?thesis proof (cases "fst instr = arith_type ADD \ fst instr = arith_type ADDcc") case True then show ?thesis using f1 f3 apply clarsimp proof - assume a1: "low_equal s1 s2" assume "t1 = snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1))))))))" assume "t2 = snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))))" have f2: "\is. get_operand2 is s1 = get_operand2 is s2" using a1 by (meson get_operand2_low_equal) have f3: "fst (fst (get_curr_win () s1)) = fst (fst (get_curr_win () s2))" using a1 by (meson get_curr_win_low_equal) have "\w wa. user_reg_val wa w (snd (fst (get_curr_win () s1))) = user_reg_val wa w (snd (fst (get_curr_win () s2)))" using a1 by (meson get_curr_win2_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))" using f3 f2 a1 by (metis (no_types) add_instr_sub1_low_equal get_curr_win2_low_equal write_reg_low_equal) qed next case False then show ?thesis using f1 f3 apply clarsimp proof - assume a1: "low_equal s1 s2" have f2: "gets (\s. ucast (get_CWP (cpu_reg_val PSR s))::'a word) = get_curr_win ()" by (simp add: get_curr_win_def) then have "((ucast (get_CWP (cpu_reg_val PSR s1)), s1), False) = (fst (get_curr_win () s1), snd (get_curr_win () s1))" by (metis (no_types) prod.collapse simpler_gets_def) then have "(ucast (get_CWP (cpu_reg_val PSR s1)), s1) = fst (get_curr_win () s1) \ \ snd (get_curr_win () s1)" by fastforce then have f3: "ucast (get_CWP (cpu_reg_val PSR s1)) = fst (fst (get_curr_win () s1)) \ s1 = snd (fst (get_curr_win () s1))" by (metis prod.collapse prod.simps(1)) have "((ucast (get_CWP (cpu_reg_val PSR s2)), s2), False) = (fst (get_curr_win () s2), snd (get_curr_win () s2))" using f2 by (metis (no_types) prod.collapse simpler_gets_def) then have "(ucast (get_CWP (cpu_reg_val PSR s2)), s2) = fst (get_curr_win () s2) \ \ snd (get_curr_win () s2)" by fastforce then have f4: "ucast (get_CWP (cpu_reg_val PSR s2)) = fst (fst (get_curr_win () s2)) \ s2 = snd (fst (get_curr_win () s2))" by (metis (no_types) prod.collapse prod.simps(1)) then have f5: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2))) = low_equal s1 s2" using f3 by presburger have f6: "fst (fst (get_curr_win () s1)) = ucast (get_CWP (cpu_reg_val PSR s1))" using f3 by auto have f7: "cpu_reg_val PSR s1 = cpu_reg_val PSR s2" using a1 by (meson cpu_reg_val_low_equal) have f8: "\s sa w wa. \ low_equal s sa \ user_reg_val (w::'a word) wa s = user_reg_val w wa sa" by (meson user_reg_val_low_equal) have f9: "ucast (get_CWP (cpu_reg_val PSR s2)) = fst (fst (get_curr_win () s2))" using f4 by meson have "\s sa is. \ low_equal (s::'a sparc_state) sa \ get_operand2 is s = get_operand2 is sa" using get_operand2_low_equal by blast then have f10: "get_operand2 (snd instr) s1 = get_operand2 (snd instr) s2" using a1 by meson have f11: "cpu_reg_val PSR (snd (fst (get_curr_win () s2))) = cpu_reg_val PSR s1" using f4 a1 by (simp add: cpu_reg_val_low_equal) have f12: "user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1))) = 0" by (meson user_reg_val_def) have "user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))) = 0" by (meson user_reg_val_def) then have "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))" using f12 f9 f7 f5 f3 a1 write_reg_low_equal by fastforce then have "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))) \ snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))))))) = snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (if get_operand_w5 (snd instr ! Suc 0) = 0 then 0 else user_reg (snd (fst (get_curr_win () s2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))))))) \ snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))) = snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (if get_operand_w5 (snd instr ! Suc 0) = 0 then 0 else user_reg (snd (fst (get_curr_win () s2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))))" using f11 f10 f9 f8 f7 f6 f5 f3 a1 by (simp add: user_reg_val_def) then show "low_equal (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))" using add_instr_sub1_low_equal by blast qed qed qed qed lemma sub_instr_sub1_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (sub_instr_sub1 instr_name result rs1_val operand2 s1)) \ t2 = snd (fst (sub_instr_sub1 instr_name result rs1_val operand2 s2))" shows "low_equal t1 t2" proof (cases "instr_name = arith_type SUBcc \ instr_name = arith_type SUBXcc") case True then show ?thesis using a1 apply (simp add: sub_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (clarsimp simp add: cpu_reg_val_low_equal) using write_cpu_low_equal by blast next case False then show ?thesis using a1 apply (simp add: sub_instr_sub1_def) by (simp add: return_def) qed lemma sub_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (sub_instr instr s1)) \ t2 = snd (fst (sub_instr instr s2))" shows "low_equal t1 t2" proof - have f1: "low_equal s1 s2 \ t1 = snd (fst (sub_instr_sub1 (fst instr) (if fst instr = arith_type SUB \ fst instr = arith_type SUBcc then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 else user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) \ 0 then if fst instr = arith_type SUB \ fst instr = arith_type SUBcc then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 else user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1))))) else user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))))) \ t2 = snd (fst (sub_instr_sub1 (fst instr) (if fst instr = arith_type SUB \ fst instr = arith_type SUBcc then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 else user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) \ 0 then if fst instr = arith_type SUB \ fst instr = arith_type SUBcc then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 else user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2))))) else user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))))" using a1 apply (simp add: sub_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: case_prod_unfold) then show ?thesis proof (cases "get_operand_w5 (snd instr ! 3) \ 0") case True then have f2: "get_operand_w5 (snd instr ! 3) \ 0" by auto then show ?thesis proof (cases "fst instr = arith_type SUB \ fst instr = arith_type SUBcc") case True then show ?thesis using f1 f2 apply clarsimp proof - assume a1: "low_equal s1 s2" assume a2: "t1 = snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))" assume a3: "t2 = snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))))" then have f4: "snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s1) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))) = t2" using a1 by (simp add: get_operand2_low_equal) have "\s. \ low_equal (snd (fst (get_curr_win () s1))) s \ snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) s - get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) s) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) s - get_operand2 (snd instr) s1) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))))) = t1" using a2 a1 by (simp add: get_curr_win_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))) (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))))" using f4 a3 a2 a1 by (metis (no_types) get_curr_win2_low_equal sub_instr_sub1_low_equal write_reg_low_equal) qed next case False then show ?thesis using f1 f2 apply clarsimp proof - assume a1: "low_equal s1 s2" have f2: "fst (get_curr_win () s1) = (ucast (get_CWP (cpu_reg_val PSR s1)), s1)" by (simp add: get_curr_win_def simpler_gets_def) have f3: "cpu_reg_val PSR s1 = cpu_reg_val PSR s2" using a1 by (meson cpu_reg_val_low_equal) then have f4: "user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) = user_reg_val (ucast (get_CWP (cpu_reg_val PSR s2))) (get_operand_w5 (snd instr ! Suc 0)) s1" using f2 by simp have f5: "\s sa is. \ low_equal (s::'a sparc_state) sa \ get_operand2 is s = get_operand2 is sa" using get_operand2_low_equal by blast then have f6: "sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (ucast (get_CWP (cpu_reg_val PSR s2))) (get_operand_w5 (snd instr ! Suc 0)) s2) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) = sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))" using f4 a1 by (simp add: user_reg_val_low_equal) have f7: "fst (get_curr_win () s2) = (ucast (get_CWP (cpu_reg_val PSR s2)), s2)" by (simp add: get_curr_win_def simpler_gets_def) then have f8: "user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2))))) = user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))" using f5 f2 a1 by (simp add: cpu_reg_val_low_equal user_reg_val_low_equal) then have f9: "sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (ucast (get_CWP (cpu_reg_val PSR s2))) (get_operand_w5 (snd instr ! Suc 0)) s2) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))) = sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using f7 by fastforce have "write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (ucast (get_CWP (cpu_reg_val PSR s2))) (get_operand_w5 (snd instr ! 3)) s2 = write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))" using f8 f7 by simp then have "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using f3 f2 a1 by (metis (no_types) prod.sel(1) prod.sel(2) write_reg_low_equal) then show "low_equal (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))) (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))))" using f9 f6 by (metis (no_types) sub_instr_sub1_low_equal) qed qed next case False then have f3: "\ get_operand_w5 (snd instr ! 3) \ 0" by auto then show ?thesis proof (cases "fst instr = arith_type SUB \ fst instr = arith_type SUBcc") case True then show ?thesis using f1 f3 apply clarsimp proof - assume a1: "low_equal s1 s2" assume "t1 = snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1))))))))" assume "t2 = snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))))" have f2: "\is. get_operand2 is s1 = get_operand2 is s2" using a1 get_operand2_low_equal by blast have f3: "fst (fst (get_curr_win () s1)) = fst (fst (get_curr_win () s2))" using a1 by (meson get_curr_win_low_equal) have "\w wa. user_reg_val wa w (snd (fst (get_curr_win () s1))) = user_reg_val wa w (snd (fst (get_curr_win () s2)))" using a1 by (metis (no_types) get_curr_win2_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))" using f3 f2 a1 by (metis (no_types) get_curr_win2_low_equal sub_instr_sub1_low_equal write_reg_low_equal) qed next case False then show ?thesis using f1 f3 apply clarsimp proof - assume a1: "low_equal s1 s2" have f2: "\s sa sb w wa wb sc. (\ low_equal s sa \ sb \ snd (fst (write_reg w (wa::'a word) wb s)) \ sc \ snd (fst (write_reg w wa wb sa))) \ low_equal sb sc" by (meson write_reg_low_equal) have "((ucast (get_CWP (cpu_reg_val PSR s1)), s1), False) = get_curr_win () s1" by (simp add: get_curr_win_def simpler_gets_def) then have f3: "ucast (get_CWP (cpu_reg_val PSR s1)) = fst (fst (get_curr_win () s1)) \ s1 = snd (fst (get_curr_win () s1))" by (metis (no_types) prod.collapse prod.simps(1)) have "((ucast (get_CWP (cpu_reg_val PSR s2)), s2), False) = get_curr_win () s2" by (simp add: get_curr_win_def simpler_gets_def) then have f4: "ucast (get_CWP (cpu_reg_val PSR s2)) = fst (fst (get_curr_win () s2)) \ s2 = snd (fst (get_curr_win () s2))" by (metis (no_types) prod.collapse prod.simps(1)) have f5: "\s sa sb sc w wa wb sd. (\ low_equal (s::'a sparc_state) sa \ sb \ snd (fst (sub_instr_sub1 sc w wa wb s)) \ sd \ snd (fst (sub_instr_sub1 sc w wa wb sa))) \ low_equal sb sd" by (meson sub_instr_sub1_low_equal) have "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))" using f4 f3 f2 a1 by (simp add: cpu_reg_val_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))" using f5 f4 f3 a1 by (simp add: cpu_reg_val_low_equal get_operand2_low_equal user_reg_val_low_equal) qed qed qed qed lemma mul_instr_sub1_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (mul_instr_sub1 instr_name result s1)) \ t2 = snd (fst (mul_instr_sub1 instr_name result s2))" shows "low_equal t1 t2" proof (cases "instr_name \ {arith_type SMULcc,arith_type UMULcc}") case True then show ?thesis using a1 apply (simp add: mul_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (clarsimp simp add: cpu_reg_val_low_equal) using write_cpu_low_equal by blast next case False then show ?thesis using a1 apply (simp add: mul_instr_sub1_def) by (simp add: return_def) qed lemma mul_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (mul_instr instr s1)) \ t2 = snd (fst (mul_instr instr s2))" shows "low_equal t1 t2" using a1 apply (simp add: mul_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) proof - assume a1: "low_equal s1 s2 \ t1 = snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) \ 0 then ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) else user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))))))) \ t2 = snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) \ 0 then ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) else user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))" have f2: "\s sa sb sc w sd. \ low_equal (s::'a sparc_state) sa \ sb \ snd (fst (mul_instr_sub1 sc w s)) \ sd \ snd (fst (mul_instr_sub1 sc w sa)) \ low_equal sb sd" using mul_instr_sub1_low_equal by blast have f3: "\s sa sb w wa wb sc. \ low_equal s sa \ sb \ snd (fst (write_reg w (wa::'a word) wb s)) \ sc \ snd (fst (write_reg w wa wb sa)) \ low_equal sb sc" by (meson write_reg_low_equal) have f4: "\s sa sb w c sc. \ low_equal (s::'a sparc_state) sa \ sb \ snd (fst (write_cpu w c s)) \ sc \ snd (fst (write_cpu w c sa)) \ low_equal sb sc" by (meson write_cpu_low_equal) have f5: "low_equal s1 s2 \ t1 = snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))))))) \ t2 = snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))" using a1 by presburger have f6: "((ucast (get_CWP (cpu_reg_val PSR s1)), s1), False) = (fst (get_curr_win () s1), snd (get_curr_win () s1))" by (simp add: get_curr_win_def simpler_gets_def) have f7: "fst (fst (get_curr_win () s1)) = fst (fst (get_curr_win () s2))" using f5 by (meson get_curr_win_low_equal) have "((ucast (get_CWP (cpu_reg_val PSR s2)), s2), False) = (fst (get_curr_win () s2), snd (get_curr_win () s2))" by (simp add: get_curr_win_def simpler_gets_def) then have f8: "ucast (get_CWP (cpu_reg_val PSR s2)) = fst (fst (get_curr_win () s2)) \ s2 = snd (fst (get_curr_win () s2))" by (metis prod.collapse prod.simps(1)) then have f9: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using f6 f5 by (metis (no_types) prod.collapse prod.simps(1)) have f10: "\s sa w wa. \ low_equal s sa \ user_reg_val (w::'a word) wa s = user_reg_val w wa sa" using user_reg_val_low_equal by blast have f11: "get_operand2 (snd instr) s1 = get_operand2 (snd instr) (snd (fst (get_curr_win () s2)))" using f9 f6 by (metis (no_types) get_operand2_low_equal prod.collapse prod.simps(1)) then have f12: "uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2) = uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)" using f10 f9 f8 f7 by presburger then have f13: "(word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) \ (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) \ (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) \ (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) \ low_equal (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))" using f9 f4 by presburger have "get_operand_w5 (snd instr ! 3) = 0 \ low_equal (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) = write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))" using f10 f7 by force then have f14: "get_operand_w5 (snd instr ! 3) \ 0 \ low_equal (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))))) \ \ low_equal (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))" using f3 by metis then have f15: "low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ fst instr \ arith_type UMULcc \ get_operand_w5 (snd instr ! 3) \ 0 \ (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) \ (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) \ (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) \ (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))" using f13 f12 f2 by fastforce have f16: "user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) = user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))" using f10 f9 f7 by presburger { assume "fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))" moreover { assume "\ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))" moreover { assume "low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))" moreover { assume "mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))))) \ mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))" then have "write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) \ fst instr \ arith_type UMULcc" by fastforce } moreover { assume "mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))) \ mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))" then have "write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ fst instr \ arith_type UMULcc" by fastforce } ultimately have "write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) \ fst instr \ arith_type UMULcc" by force } ultimately have "write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) \ \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ fst instr \ arith_type UMULcc" by fastforce } ultimately have "fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))) = write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) \ write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))) = write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))" by blast } moreover { assume "\ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))" moreover { assume "\ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))))))) = snd (fst (mul_instr_sub1 (arith_type UMULcc) (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))))))) \ snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))) = snd (fst (mul_instr_sub1 (arith_type UMULcc) (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))" then have "\ low_equal (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))" using f2 by blast moreover { assume "(if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) \ (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))))" moreover { assume "(if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) \ ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))" then have "(if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) = (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) \ get_operand_w5 (snd instr ! 3) = 0" by (metis f11 f16 f8) } ultimately have "(if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) = (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) \ get_operand_w5 (snd instr ! 3) = 0" by fastforce } ultimately have "fst instr = arith_type UMULcc \ (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) = (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) \ get_operand_w5 (snd instr ! 3) = 0" using f13 f7 f3 by fastforce } moreover { assume "mul_instr_sub1 (arith_type UMULcc) (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))))) \ mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))" moreover { assume "(if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) \ ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))" then have "(if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) = (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) \ get_operand_w5 (snd instr ! 3) = 0" by (metis f11 f16 f8) } ultimately have "fst instr = arith_type UMULcc \ (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) = (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) \ get_operand_w5 (snd instr ! 3) = 0" by fastforce } ultimately have "fst instr = arith_type UMULcc \ get_operand_w5 (snd instr ! 3) = 0" using f12 by fastforce } moreover { assume "write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))" then have "fst instr = arith_type UMULcc \ get_operand_w5 (snd instr ! 3) = 0" by presburger } ultimately have "fst instr = arith_type UMULcc \ get_operand_w5 (snd instr ! 3) = 0 \ get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))" by force moreover { assume "fst instr \ arith_type UMULcc" { assume "fst instr \ arith_type UMULcc \ low_equal (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))" moreover { assume "fst instr \ arith_type UMULcc \ low_equal (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))))) \ snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))) = snd (fst (mul_instr_sub1 (arith_type UMUL) (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))" then have "(fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ fst instr \ arith_type UMUL \ fst instr \ arith_type UMULcc \ (get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ get_operand_w5 (snd instr ! 3) = 0" using f2 by presburger } ultimately have "fst instr \ arith_type UMULcc \ (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) \ ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ fst instr \ arith_type UMUL \ fst instr \ arith_type UMULcc \ (get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ get_operand_w5 (snd instr ! 3) = 0" by fastforce } then have "(get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ get_operand_w5 (snd instr ! 3) = 0" using f16 f11 f9 f8 f7 f4 f3 f2 by force } moreover { assume "get_operand_w5 (snd instr ! 3) = 0" moreover { assume "(fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0" moreover { assume "((fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0) \ \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))" moreover { assume "((fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0) \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))" then have "((fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0) \ arith_type UMUL \ arith_type UMULcc \ fst instr \ arith_type UMULcc \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0" by fastforce then have "((fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0) \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ fst instr \ arith_type UMULcc \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0" by force } ultimately have "((fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0) \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ ((fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0) \ \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ fst instr \ arith_type UMULcc \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0" by simp } ultimately have "((fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0) \ \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ fst instr \ arith_type UMULcc \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0 \ (get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ (get_operand_w5 (snd instr ! 3) \ 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))))" by auto then have "fst instr \ arith_type UMULcc \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0 \ (get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ (get_operand_w5 (snd instr ! 3) \ 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))))" using f15 by presburger then have "(get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ (get_operand_w5 (snd instr ! 3) \ 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))))" using f14 f13 f12 f2 by force } ultimately have "(get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ (get_operand_w5 (snd instr ! 3) \ 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))))" using f16 f14 f11 f9 f8 f4 f2 by fastforce } ultimately show "(get_operand_w5 (snd instr ! 3) \ 0 \ (fst instr = arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMUL \ fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ (get_operand_w5 (snd instr ! 3) = 0 \ (fst instr = arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMUL \ fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))))" by blast qed lemma div_write_new_val_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (div_write_new_val i result temp_V s1)) \ t2 = snd (fst (div_write_new_val i result temp_V s2))" shows "low_equal t1 t2" proof (cases "(fst i) \ {arith_type UDIVcc,arith_type SDIVcc}") case True then show ?thesis using a1 apply (simp add: div_write_new_val_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (clarsimp simp add: cpu_reg_val_low_equal) using write_cpu_low_equal by blast next case False then show ?thesis using a1 apply (simp add: div_write_new_val_def) by (simp add: return_def) qed lemma div_comp_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (div_comp instr rs1 rd operand2 s1)) \ t2 = snd (fst (div_comp instr rs1 rd operand2 s2))" shows "low_equal t1 t2" using a1 apply (simp add: div_comp_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (clarsimp simp add: get_curr_win_low_equal) proof - assume a1: "low_equal s1 s2" have f2: "\s sa sb w wa wb sc. \ low_equal s sa \ sb \ snd (fst (write_reg w (wa::'a word) wb s)) \ sc \ snd (fst (write_reg w wa wb sa)) \ low_equal sb sc" by (meson write_reg_low_equal) have f3: "gets (\s. ucast (get_CWP (cpu_reg_val PSR s))::'a word) = get_curr_win ()" by (simp add: get_curr_win_def) then have "((ucast (get_CWP (cpu_reg_val PSR s1)), s1), False) = (fst (get_curr_win () s1), snd (get_curr_win () s1))" by (metis (no_types) prod.collapse simpler_gets_def) then have f4: "ucast (get_CWP (cpu_reg_val PSR s1)) = fst (fst (get_curr_win () s1)) \ s1 = snd (fst (get_curr_win () s1))" by (metis prod.collapse prod.simps(1)) have "((ucast (get_CWP (cpu_reg_val PSR s2)), s2), False) = (fst (get_curr_win () s2), snd (get_curr_win () s2))" using f3 by (metis (no_types) prod.collapse simpler_gets_def) then have f5: "ucast (get_CWP (cpu_reg_val PSR s2)) = fst (fst (get_curr_win () s2)) \ s2 = snd (fst (get_curr_win () s2))" by (metis (no_types) prod.collapse prod.simps(1)) then have f6: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using f4 a1 by presburger have f7: "\s sa sb p w wa sc. \ low_equal (s::'a sparc_state) sa \ sb \ snd (fst (div_write_new_val p w wa s)) \ sc \ snd (fst (div_write_new_val p w wa sa)) \ low_equal sb sc" by (meson div_write_new_val_low_equal) have f8: "cpu_reg_val PSR s2 = cpu_reg_val PSR s1" using a1 by (simp add: cpu_reg_val_def low_equal_def) then have "fst (fst (get_curr_win () s2)) = ucast (get_CWP (cpu_reg_val PSR s1))" using f5 by presburger then have f9: "fst (fst (get_curr_win () s2)) = fst (fst (get_curr_win () s1))" using f4 by presburger have f10: "fst (fst (get_curr_win () s1)) = fst (fst (get_curr_win () s2))" using f8 f5 f4 by presburger have f11: "(word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))::word64) = word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))" using f5 f4 a1 by (metis (no_types) cpu_reg_val_def low_equal_def user_reg_val_low_equal) have f12: "ucast (get_CWP (cpu_reg_val PSR s1)) = fst (fst (get_curr_win () s2))" using f8 f5 by presburger then have "rd = 0 \ (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) = user_reg_val (ucast (get_CWP (cpu_reg_val PSR s1))) 0 (snd (fst (get_curr_win () s1)))" using f6 user_reg_val_low_equal by fastforce then have f13: "rd = 0 \ write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (ucast (get_CWP (cpu_reg_val PSR s1))) 0 (snd (fst (get_curr_win () s1))) = write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1)))" using f12 f10 by presburger have f14: "write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (ucast (get_CWP (cpu_reg_val PSR s1))) rd (snd (fst (get_curr_win () s2))) = write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))" using f12 f11 by auto have "write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (ucast (get_CWP (cpu_reg_val PSR s1))) rd (snd (fst (get_curr_win () s1))) = write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))) \ write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (ucast (get_CWP (cpu_reg_val PSR s1))) rd (snd (fst (get_curr_win () s2))) = write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) \ low_equal (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))))))" using f6 f2 by metis moreover { assume "low_equal (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))))))" then have "low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))))))))" using f11 f9 f7 by metis moreover { assume "low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))))))))) \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))))))))" then have "div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2) = (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) \ rd = 0" by fastforce } ultimately have "div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2) = (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) \ rd = 0 \ (rd \ 0 \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))))))))) \ (rd = 0 \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))))))" by fastforce } moreover { assume "write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (ucast (get_CWP (cpu_reg_val PSR s1))) rd (snd (fst (get_curr_win () s1))) \ write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1)))" then have "div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2) \ (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2))" using f12 f9 by fastforce } moreover { assume "write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (ucast (get_CWP (cpu_reg_val PSR s1))) rd (snd (fst (get_curr_win () s2))) \ write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))" then have "rd = 0" using f14 by presburger } moreover { assume "rd = 0" then have "rd = 0 \ low_equal (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))))))" using f13 f12 f6 f2 by metis then have "rd = 0 \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))))))))" using f11 f9 f7 by metis then have "(rd \ 0 \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))))))))) \ (rd = 0 \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))))))" using f10 by fastforce } ultimately show "(rd \ 0 \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))))))))) \ (rd = 0 \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))))))" using f9 by fastforce qed lemma div_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (div_instr instr s1)) \ t2 = snd (fst (div_instr instr s2))" shows "low_equal t1 t2" using a1 apply (simp add: div_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: return_def) apply (auto simp add: get_operand2_low_equal) apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (auto simp add: traps_low_equal) apply (blast intro: mod_trap_low_equal) using div_comp_low_equal by blast lemma get_curr_win_traps_low_equal: assumes a1: "low_equal s1 s2" shows "low_equal (snd (fst (get_curr_win () s1)) \traps := insert some_trap (traps (snd (fst (get_curr_win () s1))))\) (snd (fst (get_curr_win () s2)) \traps := insert some_trap (traps (snd (fst (get_curr_win () s2))))\)" proof - from a1 have f1: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using get_curr_win2_low_equal by auto then have f2: "(traps (snd (fst (get_curr_win () s1)))) = (traps (snd (fst (get_curr_win () s2))))" using traps_low_equal by auto then show ?thesis using f1 f2 mod_trap_low_equal by fastforce qed lemma save_restore_instr_sub1_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (save_retore_sub1 result new_cwp rd s1)) \ t2 = snd (fst (save_retore_sub1 result new_cwp rd s2))" shows "low_equal t1 t2" using a1 apply (simp add: save_retore_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (auto simp add: cpu_reg_val_low_equal) using write_cpu_low_equal write_reg_low_equal by fastforce lemma get_WIM_bit_low_equal: assumes a1: "low_equal s1 s2" shows "get_WIM_bit (unat (((word_of_int ((uint (fst (fst (get_curr_win () s1))) - 1) mod NWINDOWS)))::word5)) (cpu_reg_val WIM (snd (fst (get_curr_win () s1)))) = get_WIM_bit (unat (((word_of_int ((uint (fst (fst (get_curr_win () s2))) -1) mod NWINDOWS)))::word5)) (cpu_reg_val WIM (snd (fst (get_curr_win () s2))))" proof - from a1 have f1: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using get_curr_win2_low_equal by blast then have f2: "(cpu_reg_val WIM (snd (fst (get_curr_win () s1)))) = (cpu_reg_val WIM (snd (fst (get_curr_win () s2))))" using cpu_reg_val_low_equal by auto from a1 have "(fst (fst (get_curr_win () s1))) = (fst (fst (get_curr_win () s2)))" using get_curr_win_low_equal by auto then show ?thesis using f1 f2 by auto qed lemma get_WIM_bit_low_equal2: assumes a1: "low_equal s1 s2" shows "get_WIM_bit (unat (((word_of_int ((uint (fst (fst (get_curr_win () s1))) + 1) mod NWINDOWS)))::word5)) (cpu_reg_val WIM (snd (fst (get_curr_win () s1)))) = get_WIM_bit (unat (((word_of_int ((uint (fst (fst (get_curr_win () s2))) + 1) mod NWINDOWS)))::word5)) (cpu_reg_val WIM (snd (fst (get_curr_win () s2))))" proof - from a1 have f1: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using get_curr_win2_low_equal by blast then have f2: "(cpu_reg_val WIM (snd (fst (get_curr_win () s1)))) = (cpu_reg_val WIM (snd (fst (get_curr_win () s2))))" using cpu_reg_val_low_equal by auto from a1 have "(fst (fst (get_curr_win () s1))) = (fst (fst (get_curr_win () s2)))" using get_curr_win_low_equal by auto then show ?thesis using f1 f2 by auto qed lemma save_restore_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (save_restore_instr instr s1)) \ t2 = snd (fst (save_restore_instr instr s2))" shows "low_equal t1 t2" proof (cases "fst instr = ctrl_type SAVE") case True then have f1: "fst instr = ctrl_type SAVE" by auto then show ?thesis using a1 apply (simp add: save_restore_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply auto apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (simp add: get_curr_win_traps_low_equal) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: get_WIM_bit_low_equal) apply (simp add: get_WIM_bit_low_equal) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: get_curr_win_low_equal) using get_curr_win2_low_equal save_restore_instr_sub1_low_equal get_addr2_low_equal by metis next case False then show ?thesis using a1 apply (simp add: save_restore_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply auto apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (simp add: get_curr_win_traps_low_equal) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: get_WIM_bit_low_equal2) apply (simp add: get_WIM_bit_low_equal2) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: get_curr_win_low_equal) using get_curr_win2_low_equal save_restore_instr_sub1_low_equal get_addr2_low_equal by metis qed lemma call_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (call_instr instr s1)) \ t2 = snd (fst (call_instr instr s2))" shows "low_equal t1 t2" using a1 apply (simp add: call_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (auto simp add: get_curr_win_low_equal) using cpu_reg_val_low_equal get_curr_win2_low_equal write_cpu_low_equal write_reg_low_equal proof - assume a1: "low_equal s1 s2" assume "t1 = snd (fst (write_cpu (cpu_reg_val PC (snd (fst (get_curr_win () s1))) + (ucast (get_operand_w30 (snd instr ! 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (get_curr_win () s1)))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 15 (snd (fst (get_curr_win () s1)))))))))))" assume "t2 = snd (fst (write_cpu (cpu_reg_val PC (snd (fst (get_curr_win () s2))) + (ucast (get_operand_w30 (snd instr ! 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (get_curr_win () s2)))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 15 (snd (fst (get_curr_win () s2)))))))))))" have "\c. cpu_reg_val c (snd (fst (get_curr_win () s1))) = cpu_reg_val c (snd (fst (get_curr_win () s2)))" using a1 by (meson cpu_reg_val_low_equal get_curr_win2_low_equal) then show "low_equal (snd (fst (write_cpu (cpu_reg_val PC (snd (fst (get_curr_win () s1))) + (ucast (get_operand_w30 (snd instr ! 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (get_curr_win () s1)))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 15 (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (write_cpu (cpu_reg_val PC (snd (fst (get_curr_win () s2))) + (ucast (get_operand_w30 (snd instr ! 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (get_curr_win () s2)))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 15 (snd (fst (get_curr_win () s2))))))))))))" using a1 by (metis (no_types) get_curr_win2_low_equal write_cpu_low_equal write_reg_low_equal) qed lemma jmpl_instr_low_equal_sub1: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (write_cpu (get_addr (snd instr) (snd (fst (get_curr_win () s2)))) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))))) \ t2 = snd (fst (write_cpu (get_addr (snd instr) (snd (fst (get_curr_win () s2)))) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))))))" shows "low_equal t1 t2" proof - from a1 have f1: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using get_curr_win2_low_equal by blast then have f2: "(cpu_reg_val PC (snd (fst (get_curr_win () s1)))) = (cpu_reg_val PC (snd (fst (get_curr_win () s2))))" using cpu_reg_val_low_equal by blast then have f3: "low_equal (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using f1 write_reg_low_equal by fastforce then have "(cpu_reg_val nPC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))) = (cpu_reg_val nPC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))" using cpu_reg_val_low_equal by auto then have f4: "low_equal (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))) (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))))" using f3 write_cpu_low_equal by fastforce then show ?thesis using write_cpu_low_equal using assms by blast qed lemma jmpl_instr_low_equal_sub2: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (write_cpu (get_addr (snd instr) (snd (fst (get_curr_win () s2)))) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))) PC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))))) \ t2 = snd (fst (write_cpu (get_addr (snd instr) (snd (fst (get_curr_win () s2)))) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))) PC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))))" shows "low_equal t1 t2" proof - from a1 have f1: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using get_curr_win2_low_equal by blast then have f2: "(user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) = (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))" using user_reg_val_low_equal by blast then have f3: "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))" using f1 write_reg_low_equal by fastforce then have "(cpu_reg_val nPC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))) = (cpu_reg_val nPC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))" using cpu_reg_val_low_equal by blast then have "low_equal (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))) PC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))) PC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))" using f1 f2 f3 write_cpu_low_equal by fastforce then show ?thesis using write_cpu_low_equal using assms by blast qed lemma jmpl_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (jmpl_instr instr s1)) \ t2 = snd (fst (jmpl_instr instr s2))" shows "low_equal t1 t2" using a1 apply (simp add: jmpl_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply auto apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (simp add: get_curr_win_traps_low_equal) apply (simp add: get_addr2_low_equal) apply (simp add: get_addr2_low_equal) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp_all add: get_addr2_low_equal) apply (simp_all add: get_curr_win_low_equal) apply (case_tac "get_operand_w5 (snd instr ! 3) \ 0") apply auto using jmpl_instr_low_equal_sub1 apply blast apply (simp_all add: get_curr_win_low_equal) using jmpl_instr_low_equal_sub2 by blast lemma rett_instr_low_equal: assumes a1: "low_equal s1 s2 \ \ snd (rett_instr instr s1) \ \ snd (rett_instr instr s2) \ ((ucast (get_S (cpu_reg_val PSR s1)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s2)))::word1) = 0 \ t1 = snd (fst (rett_instr instr s1)) \ t2 = snd (fst (rett_instr instr s2))" shows "low_equal t1 t2" using a1 apply (simp add: rett_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply auto apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (simp add: return_def) using mod_trap_low_equal traps_low_equal apply fastforce using cpu_reg_val_low_equal apply fastforce using cpu_reg_val_low_equal apply fastforce apply (simp add: bind_def h1_def h2_def Let_def) by (simp add: case_prod_unfold fail_def) lemma read_state_reg_low_equal: assumes a1: "low_equal s1 s2 \ ((ucast (get_S (cpu_reg_val PSR s1)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s2)))::word1) = 0 \ t1 = snd (fst (read_state_reg_instr instr s1)) \ t2 = snd (fst (read_state_reg_instr instr s2))" shows "low_equal t1 t2" proof (cases "(fst instr \ {sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR} \ (fst instr = sreg_type RDASR \ privileged_ASR (get_operand_w5 ((snd instr)!0))))") case True then have "(fst instr \ {sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR} \ (fst instr = sreg_type RDASR \ privileged_ASR (get_operand_w5 ((snd instr)!0)))) \ ((ucast (get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s1))))))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s2))))))::word1) = 0" by (metis assms get_curr_win_privilege) then show ?thesis using a1 apply (simp add: read_state_reg_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply clarsimp using get_curr_win_traps_low_equal by auto next case False then have f1: "\((fst instr = sreg_type RDPSR \ fst instr = sreg_type RDWIM \ fst instr = sreg_type RDTBR \ fst instr = sreg_type RDASR \ privileged_ASR (get_operand_w5 (snd instr ! 0))))" by blast then show ?thesis proof (cases "illegal_instruction_ASR (get_operand_w5 ((snd instr)!0))") case True then show ?thesis using a1 f1 apply read_state_reg_instr_privilege_proof by (simp add: illegal_instruction_ASR_def) next case False then have f2: "\(illegal_instruction_ASR (get_operand_w5 ((snd instr)!0)))" by auto then show ?thesis proof (cases "(get_operand_w5 ((snd instr)!1)) \ 0") case True then have f3: "(get_operand_w5 ((snd instr)!1)) \ 0" by auto then show ?thesis proof (cases "fst instr = sreg_type RDY") case True then show ?thesis using a1 f1 f2 f3 apply (simp add: read_state_reg_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (auto simp add: get_curr_win_low_equal) using cpu_reg_val_low_equal get_curr_win2_low_equal write_reg_low_equal proof - assume "low_equal s1 s2" then have "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" by (meson get_curr_win2_low_equal) then show "low_equal (snd (fst (write_reg (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))))))" using cpu_reg_val_low_equal write_reg_low_equal by fastforce qed next case False then have f4: "\(fst instr = sreg_type RDY)" by auto then show ?thesis proof (cases "fst instr = sreg_type RDASR") case True then show ?thesis using a1 f1 f2 f3 f4 apply read_state_reg_instr_privilege_proof apply (clarsimp simp add: get_curr_win_low_equal) using cpu_reg_val_low_equal get_curr_win2_low_equal write_reg_low_equal proof - assume a1: "low_equal s1 s2" then have "cpu_reg_val (ASR (get_operand_w5 (snd instr ! 0))) (snd (fst (get_curr_win () s1))) = cpu_reg_val (ASR (get_operand_w5 (snd instr ! 0))) (snd (fst (get_curr_win () s2)))" by (meson cpu_reg_val_low_equal get_curr_win2_low_equal) then show "low_equal (snd (fst (write_reg (cpu_reg_val (ASR (get_operand_w5 (snd instr ! 0))) (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (cpu_reg_val (ASR (get_operand_w5 (snd instr ! 0))) (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))))))" using a1 by (metis (no_types) get_curr_win2_low_equal write_reg_low_equal) qed next case False then have f5: "\(fst instr = sreg_type RDASR)" by auto then show ?thesis using a1 f1 f2 f3 f4 f5 apply read_state_reg_instr_privilege_proof apply (clarsimp simp add: get_curr_win_low_equal) using cpu_reg_val_low_equal get_curr_win2_low_equal write_reg_low_equal proof - assume a1: "low_equal s1 s2" assume a2: "t1 = snd (fst (write_reg (cpu_reg_val TBR (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))))" assume "t2 = snd (fst (write_reg (cpu_reg_val TBR (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))))" have "\s. \ low_equal (snd (fst (get_curr_win () s1))) s \ snd (fst (write_reg (cpu_reg_val TBR s) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))))) = t1" using a2 by (simp add: cpu_reg_val_low_equal) then show "low_equal (snd (fst (write_reg (cpu_reg_val TBR (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (cpu_reg_val TBR (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))))))" using a2 a1 by (metis (no_types) get_curr_win2_low_equal write_reg_low_equal) qed qed qed next case False then show ?thesis using a1 f1 f2 apply (simp add: read_state_reg_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: return_def) apply clarsimp apply (simp add: case_prod_unfold) using get_curr_win2_low_equal by auto qed qed qed lemma get_s_get_curr_win: assumes a1: "low_equal s1 s2" shows "get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))) = get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s2))))" proof - from a1 have "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using get_curr_win2_low_equal by blast then show ?thesis using cpu_reg_val_low_equal by fastforce qed lemma write_state_reg_low_equal: assumes a1: "low_equal s1 s2 \ ((ucast (get_S (cpu_reg_val PSR s1)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s2)))::word1) = 0 \ t1 = snd (fst (write_state_reg_instr instr s1)) \ t2 = snd (fst (write_state_reg_instr instr s2))" shows "low_equal t1 t2" proof (cases "fst instr = sreg_type WRY") case True then show ?thesis using a1 apply write_state_reg_instr_privilege_proof apply (simp add: simpler_modify_def) apply (simp add: delayed_pool_add_def DELAYNUM_def) apply (auto simp add: get_curr_win_low_equal) using get_curr_win2_low_equal cpu_reg_mod_low_equal user_reg_val_low_equal get_operand2_low_equal proof - assume a1: "low_equal s1 s2" assume "t2 = cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s2)))) Y (snd (fst (get_curr_win () s2)))" assume "t1 = cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s1)))) Y (snd (fst (get_curr_win () s1)))" have f2: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using a1 by (meson get_curr_win2_low_equal) then have f3: "\w wa. user_reg_val w wa (snd (fst (get_curr_win () s2))) = user_reg_val w wa (snd (fst (get_curr_win () s1)))" by (simp add: user_reg_val_low_equal) have "\is. get_operand2 is (snd (fst (get_curr_win () s2))) = get_operand2 is (snd (fst (get_curr_win () s1)))" using f2 by (simp add: get_operand2_low_equal) then show "low_equal (cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s1)))) Y (snd (fst (get_curr_win () s1)))) (cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s2)))) Y (snd (fst (get_curr_win () s2))))" using f3 f2 by (metis cpu_reg_mod_low_equal) qed next case False then have f1: "\(fst instr = sreg_type WRY)" by auto then show ?thesis proof (cases "fst instr = sreg_type WRASR") case True then have f1_1: "fst instr = sreg_type WRASR" by auto then show ?thesis proof (cases "privileged_ASR (get_operand_w5 (snd instr ! 3)) \ get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))) = 0") case True then show ?thesis using a1 f1 f1_1 apply write_state_reg_instr_privilege_proof apply (clarsimp simp add: get_s_get_curr_win) apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (clarsimp simp add: get_curr_win3_low_equal) using traps_low_equal mod_trap_low_equal get_curr_win2_low_equal by fastforce next case False then have f1_2: "\ (privileged_ASR (get_operand_w5 (snd instr ! 3)) \ get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))) = 0)" by auto then show ?thesis proof (cases "illegal_instruction_ASR (get_operand_w5 (snd instr ! 3))") case True then show ?thesis using a1 f1 f1_1 f1_2 apply write_state_reg_instr_privilege_proof apply (clarsimp simp add: get_s_get_curr_win) apply auto apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (clarsimp simp add: get_curr_win3_low_equal) using traps_low_equal mod_trap_low_equal get_curr_win2_low_equal apply fastforce apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (clarsimp simp add: get_curr_win3_low_equal) using traps_low_equal mod_trap_low_equal get_curr_win2_low_equal by fastforce next case False then show ?thesis using a1 f1 f1_1 f1_2 apply write_state_reg_instr_privilege_proof apply (clarsimp simp add: get_s_get_curr_win) apply auto apply (simp add: simpler_modify_def) apply (simp add: delayed_pool_add_def DELAYNUM_def) apply (auto simp add: get_curr_win_low_equal) using get_curr_win2_low_equal cpu_reg_mod_low_equal user_reg_val_low_equal get_operand2_low_equal proof - assume a1: "low_equal s1 s2" assume "t2 = cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s2)))) (ASR (get_operand_w5 (snd instr ! 3))) (snd (fst (get_curr_win () s2)))" assume "t1 = cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s1)))) (ASR (get_operand_w5 (snd instr ! 3))) (snd (fst (get_curr_win () s1)))" have "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using a1 by (meson get_curr_win2_low_equal) then show "low_equal (cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s1)))) (ASR (get_operand_w5 (snd instr ! 3))) (snd (fst (get_curr_win () s1)))) (cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s2)))) (ASR (get_operand_w5 (snd instr ! 3))) (snd (fst (get_curr_win () s2))))" using cpu_reg_mod_low_equal get_operand2_low_equal user_reg_val_low_equal by fastforce next assume f1: "\ illegal_instruction_ASR (get_operand_w5 (snd instr ! 3))" assume f2: "fst instr = sreg_type WRASR" assume f3: "snd (fst (write_state_reg_instr instr s1)) = snd (fst (modify (delayed_pool_add (DELAYNUM, user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s1))), ASR (get_operand_w5 (snd instr ! 3)))) (snd (fst (get_curr_win () s1))))) " assume f4: "snd (fst (write_state_reg_instr instr s2)) = snd (fst (modify (delayed_pool_add (DELAYNUM, user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s2))), ASR (get_operand_w5 (snd instr ! 3)))) (snd (fst (get_curr_win () s2)))))" assume f5: "low_equal s1 s2" assume f6: "ucast (get_S (cpu_reg_val PSR s1)) = 0" assume f7: "ucast (get_S (cpu_reg_val PSR s2)) = 0" assume f8: "t1 = snd (fst (modify (delayed_pool_add (DELAYNUM, user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s1))), ASR (get_operand_w5 (snd instr ! 3)))) (snd (fst (get_curr_win () s1)))))" assume f9: "t2 = snd (fst (modify (delayed_pool_add (DELAYNUM, user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s2))), ASR (get_operand_w5 (snd instr ! 3)))) (snd (fst (get_curr_win () s2)))))" assume f10: "get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))) \ 0" assume f11: "(\s1 s2 t1 t2. low_equal s1 s2 \ t1 = snd (fst (get_curr_win () s1)) \ t2 = snd (fst (get_curr_win () s2)) \ low_equal t1 t2)" assume f12: "(\s1 s2 t1 w cr t2. low_equal s1 s2 \ t1 = cpu_reg_mod w cr s1 \ t2 = cpu_reg_mod w cr s2 \ low_equal t1 t2)" assume f13: "(\s1 s2 win ur. low_equal s1 s2 \ user_reg_val win ur s1 = user_reg_val win ur s2)" assume f14: "(\s1 s2 op_list. low_equal s1 s2 \ get_operand2 op_list s1 = get_operand2 op_list s2)" show "low_equal (snd (fst (modify (delayed_pool_add (DELAYNUM, user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s1))), ASR (get_operand_w5 (snd instr ! 3)))) (snd (fst (get_curr_win () s1)))))) (snd (fst (modify (delayed_pool_add (DELAYNUM, user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s2))), ASR (get_operand_w5 (snd instr ! 3)))) (snd (fst (get_curr_win () s2))))))" using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 using Sparc_Properties.ucast_0 assms get_curr_win_privilege by blast qed qed qed next case False then have f2: "\(fst instr = sreg_type WRASR)" by auto have f3: "get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))) = 0 \ get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))) = 0" using get_curr_win_privilege a1 by (metis ucast_id) then show ?thesis proof (cases "fst instr = sreg_type WRPSR") case True then show ?thesis using a1 f1 f2 f3 apply write_state_reg_instr_privilege_proof apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (clarsimp simp add: get_curr_win3_low_equal) using traps_low_equal mod_trap_low_equal get_curr_win2_low_equal by fastforce next case False then have f4: "\(fst instr = sreg_type WRPSR)" by auto then show ?thesis proof (cases "fst instr = sreg_type WRWIM") case True then show ?thesis using a1 f1 f2 f3 f4 apply write_state_reg_instr_privilege_proof apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (clarsimp simp add: get_curr_win3_low_equal) using traps_low_equal mod_trap_low_equal get_curr_win2_low_equal by fastforce next case False then have f5: "\(fst instr = sreg_type WRWIM)" by auto then show ?thesis using a1 f1 f2 f3 f4 f5 apply write_state_reg_instr_privilege_proof apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (clarsimp simp add: get_curr_win3_low_equal) using traps_low_equal mod_trap_low_equal get_curr_win2_low_equal by fastforce qed qed qed qed lemma flush_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (flush_instr instr s1)) \ t2 = snd (fst (flush_instr instr s2))" shows "low_equal t1 t2" using a1 apply (simp add: flush_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def simpler_modify_def) apply (simp add: flush_cache_all_def) apply (simp add: low_equal_def) apply (simp add: user_accessible_def) apply (simp add: mem_equal_def) by auto lemma branch_instr_sub1_low_equal: assumes a1: "low_equal s1 s2" shows "branch_instr_sub1 instr_name s1 = branch_instr_sub1 instr_name s2" using a1 apply (simp add: branch_instr_sub1_def) by (simp add: low_equal_def) lemma set_annul_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (set_annul True s1)) \ t2 = snd (fst (set_annul True s2))" shows "low_equal t1 t2" using a1 apply (simp add: set_annul_def) apply (simp add: simpler_modify_def annul_mod_def) using state_var2_low_equal state_var_low_equal by fastforce lemma branch_instr_low_equal_sub0: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (write_cpu (cpu_reg_val PC s2 + sign_ext24 (ucast (get_operand_w22 (snd instr ! Suc 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s1))))) \ t2 = snd (fst (write_cpu (cpu_reg_val PC s2 + sign_ext24 (ucast (get_operand_w22 (snd instr ! Suc 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s2)))))" shows "low_equal t1 t2" proof - from a1 have "low_equal (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s1))) (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s2)))" using write_cpu_low_equal by blast then show ?thesis using a1 write_cpu_low_equal by blast qed lemma branch_instr_low_equal_sub1: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (set_annul True (snd (fst (write_cpu (cpu_reg_val PC s2 + sign_ext24 (ucast (get_operand_w22 (snd instr ! Suc 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s1)))))))) \ t2 = snd (fst (set_annul True (snd (fst (write_cpu (cpu_reg_val PC s2 + sign_ext24 (ucast (get_operand_w22 (snd instr ! Suc 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s2))))))))" shows "low_equal t1 t2" proof - from a1 have "low_equal (snd (fst (write_cpu (cpu_reg_val PC s2 + sign_ext24 (ucast (get_operand_w22 (snd instr ! Suc 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s1)))))) (snd (fst (write_cpu (cpu_reg_val PC s2 + sign_ext24 (ucast (get_operand_w22 (snd instr ! Suc 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s2))))))" using branch_instr_low_equal_sub0 by blast then show ?thesis using a1 using set_annul_low_equal by blast qed lemma branch_instr_low_equal_sub2: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (set_annul True (snd (fst (write_cpu (cpu_reg_val nPC s2 + 4) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s1)))))))) \ t2 = snd (fst (set_annul True (snd (fst (write_cpu (cpu_reg_val nPC s2 + 4) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s2))))))))" shows "low_equal t1 t2" proof - from a1 have "low_equal (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s1))) (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s2)))" using write_cpu_low_equal by blast then have "low_equal (snd (fst (write_cpu (cpu_reg_val nPC s2 + 4) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s1)))))) (snd (fst (write_cpu (cpu_reg_val nPC s2 + 4) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s2))))))" using write_cpu_low_equal by blast then show ?thesis using a1 using set_annul_low_equal by blast qed lemma branch_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (branch_instr instr s1)) \ t2 = snd (fst (branch_instr instr s2))" shows "low_equal t1 t2" using a1 apply (simp add: branch_instr_def) apply (simp add: Let_def simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold return_def) apply clarsimp apply (simp add: branch_instr_sub1_low_equal) apply (simp_all add: cpu_reg_val_low_equal) apply (cases "branch_instr_sub1 (fst instr) s2 = 1") apply clarsimp apply (simp add: bind_def h1_def h2_def Let_def) apply (simp_all add: cpu_reg_val_low_equal) apply (simp add: case_prod_unfold) apply (cases "fst instr = bicc_type BA \ get_operand_flag (snd instr ! 0) = 1") apply clarsimp using branch_instr_low_equal_sub1 apply blast apply clarsimp apply (simp add: return_def) using branch_instr_low_equal_sub0 apply fastforce apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (cases "get_operand_flag (snd instr ! 0) = 1") apply clarsimp apply (simp_all add: cpu_reg_val_low_equal) using branch_instr_low_equal_sub2 apply metis apply (simp add: return_def) using write_cpu_low_equal by metis lemma dispath_instr_low_equal: assumes a1: "low_equal s1 s2 \ ((ucast (get_S (cpu_reg_val PSR s1)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR s2)))::word1) = 0 \ \ snd (dispatch_instruction instr s1) \ \ snd (dispatch_instruction instr s2) \ t1 = (snd (fst (dispatch_instruction instr s1))) \ t2 = (snd (fst (dispatch_instruction instr s2)))" shows "low_equal t1 t2" proof (cases "get_trap_set s1 = {}") case True then have f_no_traps: "get_trap_set s1 = {} \ get_trap_set s2 = {}" using a1 by (simp add: low_equal_def get_trap_set_def) then show ?thesis proof (cases "fst instr \ {load_store_type LDSB,load_store_type LDUB, load_store_type LDUBA,load_store_type LDUH,load_store_type LD, load_store_type LDA,load_store_type LDD}") case True then show ?thesis using a1 f_no_traps apply dispath_instr_privilege_proof by (blast intro: load_instr_low_equal) next case False then have f1: "fst instr \ {load_store_type LDSB, load_store_type LDUB, load_store_type LDUBA, load_store_type LDUH, load_store_type LD, load_store_type LDA, load_store_type LDD}" by auto then show ?thesis proof (cases "fst instr \ {load_store_type STB,load_store_type STH, load_store_type ST,load_store_type STA,load_store_type STD}") case True then show ?thesis using a1 f_no_traps f1 apply dispath_instr_privilege_proof using store_instr_low_equal by blast next case False then have f2: "\(fst instr \ {load_store_type STB,load_store_type STH, load_store_type ST,load_store_type STA,load_store_type STD})" by auto then show ?thesis proof (cases "fst instr \ {sethi_type SETHI}") case True then show ?thesis using a1 f_no_traps f1 f2 apply dispath_instr_privilege_proof by (auto intro: sethi_low_equal) next case False then have f3: "\(fst instr \ {sethi_type SETHI})" by auto then show ?thesis proof (cases "fst instr \ {nop_type NOP}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 apply dispath_instr_privilege_proof by (auto intro: nop_low_equal) next case False then have f4: "\(fst instr \ {nop_type NOP})" by auto then show ?thesis proof (cases "fst instr \ {logic_type ANDs,logic_type ANDcc,logic_type ANDN, logic_type ANDNcc,logic_type ORs,logic_type ORcc,logic_type ORN, logic_type XORs,logic_type XNOR}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 apply dispath_instr_privilege_proof using logical_instr_low_equal by blast next case False then have f5: "\(fst instr \ {logic_type ANDs,logic_type ANDcc,logic_type ANDN, logic_type ANDNcc,logic_type ORs,logic_type ORcc,logic_type ORN, logic_type XORs,logic_type XNOR})" by auto then show ?thesis proof (cases "fst instr \ {shift_type SLL,shift_type SRL,shift_type SRA}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 apply dispath_instr_privilege_proof using shift_instr_low_equal by blast next case False then have f6: "\(fst instr \ {shift_type SLL,shift_type SRL,shift_type SRA})" by auto then show ?thesis proof (cases "fst instr \ {arith_type ADD,arith_type ADDcc,arith_type ADDX}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 apply dispath_instr_privilege_proof using add_instr_low_equal by blast next case False then have f7: "\(fst instr \ {arith_type ADD,arith_type ADDcc,arith_type ADDX})" by auto then show ?thesis proof (cases "fst instr \ {arith_type SUB,arith_type SUBcc,arith_type SUBX}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 apply dispath_instr_privilege_proof using sub_instr_low_equal by blast next case False then have f8: "\(fst instr \ {arith_type SUB,arith_type SUBcc,arith_type SUBX})" by auto then show ?thesis proof (cases "fst instr \ {arith_type UMUL,arith_type SMUL,arith_type SMULcc}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 apply dispath_instr_privilege_proof using mul_instr_low_equal by blast next case False then have f9: "\(fst instr \ {arith_type UMUL,arith_type SMUL, arith_type SMULcc})" by auto then show ?thesis proof (cases "fst instr \ {arith_type UDIV,arith_type UDIVcc,arith_type SDIV}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 apply dispath_instr_privilege_proof using div_instr_low_equal by blast next case False then have f10: "\(fst instr \ {arith_type UDIV, arith_type UDIVcc,arith_type SDIV})" by auto then show ?thesis proof (cases "fst instr \ {ctrl_type SAVE,ctrl_type RESTORE}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 apply dispath_instr_privilege_proof using save_restore_instr_low_equal by blast next case False then have f11: "\(fst instr \ {ctrl_type SAVE,ctrl_type RESTORE})" by auto then show ?thesis proof (cases "fst instr \ {call_type CALL}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 apply dispath_instr_privilege_proof using call_instr_low_equal by blast next case False then have f12: "\(fst instr \ {call_type CALL})" by auto then show ?thesis proof (cases "fst instr \ {ctrl_type JMPL}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 apply dispath_instr_privilege_proof using jmpl_instr_low_equal by blast next case False then have f13: "\(fst instr \ {ctrl_type JMPL})" by auto then show ?thesis proof (cases "fst instr \ {ctrl_type RETT}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 apply dispath_instr_privilege_proof using rett_instr_low_equal by blast next case False then have f14: "\(fst instr \ {ctrl_type RETT})" by auto then show ?thesis proof (cases "fst instr \ {sreg_type RDY,sreg_type RDPSR, sreg_type RDWIM, sreg_type RDTBR}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 apply dispath_instr_privilege_proof using read_state_reg_low_equal by blast next case False then have f15: "\(fst instr \ {sreg_type RDY,sreg_type RDPSR, sreg_type RDWIM, sreg_type RDTBR})" by auto then show ?thesis proof (cases "fst instr \ {sreg_type WRY,sreg_type WRPSR, sreg_type WRWIM, sreg_type WRTBR}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 apply dispath_instr_privilege_proof using write_state_reg_low_equal by blast next case False then have f16: "\(fst instr \ {sreg_type WRY,sreg_type WRPSR, sreg_type WRWIM, sreg_type WRTBR})" by auto then show ?thesis proof (cases "fst instr \ {load_store_type FLUSH}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 apply dispath_instr_privilege_proof using flush_instr_low_equal by blast next case False then have f17: "\(fst instr \ {load_store_type FLUSH})" by auto then show ?thesis proof (cases "fst instr \ {bicc_type BE,bicc_type BNE, bicc_type BGU,bicc_type BLE,bicc_type BL,bicc_type BGE, bicc_type BNEG,bicc_type BG,bicc_type BCS,bicc_type BLEU, bicc_type BCC,bicc_type BA,bicc_type BN}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 apply dispath_instr_privilege_proof using branch_instr_low_equal by blast next case False then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 apply dispath_instr_privilege_proof by (simp add: fail_def) qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed next case False then have "get_trap_set s1 \ {} \ get_trap_set s2 \ {}" using a1 by (simp add: low_equal_def get_trap_set_def) then show ?thesis using a1 apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) by (simp add: return_def) qed lemma execute_instr_sub1_low_equal: assumes a1: "low_equal s1 s2 \ \ snd (execute_instr_sub1 instr s1) \ \ snd (execute_instr_sub1 instr s2) \ t1 = (snd (fst (execute_instr_sub1 instr s1))) \ t2 = (snd (fst (execute_instr_sub1 instr s2)))" shows "low_equal t1 t2" proof (cases "get_trap_set s1 = {}") case True then have "get_trap_set s1 = {} \ get_trap_set s2 = {}" using a1 by (simp add: low_equal_def get_trap_set_def) then show ?thesis using a1 apply (simp add: execute_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (case_tac "fst instr \ call_type CALL \ fst instr \ ctrl_type RETT \ fst instr \ ctrl_type JMPL \ fst instr \ bicc_type BE \ fst instr \ bicc_type BNE \ fst instr \ bicc_type BGU \ fst instr \ bicc_type BLE \ fst instr \ bicc_type BL \ fst instr \ bicc_type BGE \ fst instr \ bicc_type BNEG \ fst instr \ bicc_type BG \ fst instr \ bicc_type BCS \ fst instr \ bicc_type BLEU \ fst instr \ bicc_type BCC \ fst instr \ bicc_type BA \ fst instr \ bicc_type BN") apply clarsimp apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: low_equal_def) apply (simp add: cpu_reg_val_def write_cpu_def cpu_reg_mod_def) apply (simp add: simpler_modify_def return_def) apply (simp add: user_accessible_mod_cpu_reg mem_equal_mod_cpu_reg) apply clarsimp by (auto simp add: return_def) next case False then have "get_trap_set s1 \ {} \ get_trap_set s2 \ {}" using a1 by (simp add: low_equal_def get_trap_set_def) then show ?thesis using a1 apply (simp add: execute_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: return_def) qed theorem non_interference_step: assumes a1: "((ucast (get_S (cpu_reg_val PSR s1)))::word1) = 0 \ good_context s1 \ get_delayed_pool s1 = [] \ get_trap_set s1 = {} \ ((ucast (get_S (cpu_reg_val PSR s2)))::word1) = 0 \ get_delayed_pool s2 = [] \ get_trap_set s2 = {} \ good_context s2 \ low_equal s1 s2" shows "\t1 t2. Some t1 = NEXT s1 \ Some t2 = NEXT s2 \ ((ucast (get_S (cpu_reg_val PSR t1)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR t2)))::word1) = 0 \ low_equal t1 t2" proof - from a1 have "good_context s1 \ good_context s2" by auto then have "NEXT s1 = Some (snd (fst (execute_instruction () s1))) \ NEXT s2 = Some (snd (fst (execute_instruction () s2)))" by (simp add: single_step) then have "\t1 t2. Some t1 = NEXT s1 \ Some t2 = NEXT s2" by auto then have f0: "snd (execute_instruction() s1) = False \ snd (execute_instruction() s2) = False" by (auto simp add: NEXT_def case_prod_unfold) then have f1: "\t1 t2. Some t1 = NEXT s1 \ Some t2 = NEXT s2 \ ((ucast (get_S (cpu_reg_val PSR t1)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR t2)))::word1) = 0" using a1 apply (auto simp add: NEXT_def case_prod_unfold) by (auto simp add: safe_privilege) then show ?thesis proof (cases "exe_mode_val s1") case True then have f_exe0: "exe_mode_val s1" by auto then have f_exe: "exe_mode_val s1 \ exe_mode_val s2" proof - have "low_equal s1 s2" using a1 by auto then have "state_var s1 = state_var s2" by (simp add: low_equal_def) then have "exe_mode_val s1 = exe_mode_val s2" by (simp add: exe_mode_val_def) then show ?thesis using f_exe0 by auto qed then show ?thesis proof (cases "\e. fetch_instruction (delayed_pool_write s1) = Inl e") case True then have f_fetch_error: "\e. fetch_instruction (delayed_pool_write s1) = Inl e" by auto then have f_fetch_error2: "(\e. fetch_instruction (delayed_pool_write s1) = Inl e) \ (\e. fetch_instruction (delayed_pool_write s2) = Inl e)" proof - have "cpu_reg s1 = cpu_reg s2" using a1 by (simp add: low_equal_def) then have "cpu_reg_val PC s1 = cpu_reg_val PC s2" by (simp add: cpu_reg_val_def) then have "cpu_reg_val PC s1 = cpu_reg_val PC s2 \ ((ucast (get_S (cpu_reg_val PSR (delayed_pool_write s1))))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR (delayed_pool_write s2))))::word1) = 0" using a1 by (auto simp add: empty_delayed_pool_write_privilege) then show ?thesis using a1 f_fetch_error apply (simp add: fetch_instruction_def) apply (simp add: Let_def ucast_def) apply clarsimp apply (case_tac "uint (3 AND cpu_reg_val PC (delayed_pool_write s1)) = 0") apply auto apply (case_tac "fst (memory_read 8 (cpu_reg_val PC (delayed_pool_write s1)) (delayed_pool_write s1)) = None") apply auto apply (simp add: case_prod_unfold) using a1 apply (auto simp add: mem_read_delayed_write_low_equal) apply (simp add: case_prod_unfold) using a1 apply (auto simp add: mem_read_delayed_write_low_equal) apply (simp add: delayed_pool_write_def) by (simp add: Let_def get_delayed_write_def) qed then show ?thesis proof (cases "exe_mode_val s1") case True then have "exe_mode_val s1 \ exe_mode_val s2" using exe_mode_low_equal a1 by auto then show ?thesis using f1 apply (simp add: NEXT_def execute_instruction_def) apply (simp add: bind_def h1_def h2_def Let_def simpler_gets_def) using a1 apply clarsimp apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: simpler_modify_def) using f_fetch_error2 apply clarsimp apply (simp add: raise_trap_def simpler_modify_def return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: return_def simpler_modify_def) apply (simp add: raise_trap_def simpler_modify_def return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: return_def) apply (simp add: delayed_pool_write_def get_delayed_write_def Let_def) apply (simp add: low_equal_def) apply (simp add: add_trap_set_def) apply (simp add: cpu_reg_val_def) apply clarsimp by (simp add: mem_equal_mod_trap user_accessible_mod_trap) next case False then have "\ (exe_mode_val s1) \ \ (exe_mode_val s2)" using exe_mode_low_equal a1 by auto then show ?thesis using f1 apply (simp add: NEXT_def execute_instruction_def) apply (simp add: bind_def h1_def h2_def Let_def simpler_gets_def) using a1 apply clarsimp apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: return_def) qed next case False then have f_fetch_suc: "(\v. fetch_instruction (delayed_pool_write s1) = Inr v)" using fetch_instr_result_1 by auto then have "(\v. fetch_instruction (delayed_pool_write s1) = Inr v \ fetch_instruction (delayed_pool_write s2) = Inr v)" proof - have "cpu_reg s1 = cpu_reg s2" using a1 by (simp add: low_equal_def) then have "cpu_reg_val PC s1 = cpu_reg_val PC s2" by (simp add: cpu_reg_val_def) then have "cpu_reg_val PC s1 = cpu_reg_val PC s2 \ ((ucast (get_S (cpu_reg_val PSR (delayed_pool_write s1))))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR (delayed_pool_write s2))))::word1) = 0" using a1 by (auto simp add: empty_delayed_pool_write_privilege) then show ?thesis using a1 f_fetch_suc apply (simp add: fetch_instruction_def) apply (simp add: Let_def ucast_def) apply clarsimp apply (case_tac "uint (3 AND cpu_reg_val PC (delayed_pool_write s1)) = 0") apply auto apply (case_tac "fst (memory_read 8 (cpu_reg_val PC (delayed_pool_write s1)) (delayed_pool_write s1)) = None") apply auto apply (simp add: case_prod_unfold) using a1 apply (auto simp add: mem_read_delayed_write_low_equal) apply (simp add: case_prod_unfold) using a1 apply (auto simp add: mem_read_delayed_write_low_equal) apply (simp add: delayed_pool_write_def) by (simp add: Let_def get_delayed_write_def) qed then have "(\v. fetch_instruction (delayed_pool_write s1) = Inr v \ fetch_instruction (delayed_pool_write s2) = Inr v \ \ (\e. (decode_instruction v) = Inl e))" using dispatch_fail f0 a1 f_exe by auto then have f_fetch_dec: "(\v. fetch_instruction (delayed_pool_write s1) = Inr v \ fetch_instruction (delayed_pool_write s2) = Inr v \ (\v1. (decode_instruction v) = Inr v1))" using decode_instr_result_4 by auto then show ?thesis proof (cases "annul_val (delayed_pool_write s1)") case True then have "annul_val (delayed_pool_write s1) \ annul_val (delayed_pool_write s2)" using a1 apply (simp add: low_equal_def) by (simp add: delayed_pool_write_def get_delayed_write_def annul_val_def) then show ?thesis using a1 f1 f_exe f_fetch_dec apply (simp add: NEXT_def execute_instruction_def) apply (simp add: exec_gets return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: write_cpu_def cpu_reg_val_def set_annul_def) apply (simp add: simpler_modify_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: write_cpu_def cpu_reg_val_def set_annul_def) apply (simp add: simpler_modify_def) apply (simp add: cpu_reg_mod_def annul_mod_def) apply (simp add: delayed_pool_write_def get_delayed_write_def) apply (simp add: write_annul_def) apply clarsimp apply (simp add: low_equal_def) apply (simp add: user_accessible_annul mem_equal_annul) by (metis) next case False then have "\ annul_val (delayed_pool_write s1) \ \ annul_val (delayed_pool_write s2)" using a1 apply (simp add: low_equal_def) apply (simp add: delayed_pool_write_def get_delayed_write_def) by (simp add: annul_val_def) then show ?thesis using a1 f1 f_exe f_fetch_dec apply (simp add: NEXT_def execute_instruction_def) apply (simp add: exec_gets return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (case_tac "snd (execute_instr_sub1 (a, b) (snd (fst (dispatch_instruction (a, b) (delayed_pool_write s1))))) \ snd (dispatch_instruction (a, b) (delayed_pool_write s1))") apply auto apply (case_tac "snd (execute_instr_sub1 (a, b) (snd (fst (dispatch_instruction (a, b) (delayed_pool_write s2))))) \ snd (dispatch_instruction (a, b) (delayed_pool_write s2))") apply auto apply (simp add: simpler_modify_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: delayed_pool_write_def get_delayed_write_def) by (meson dispath_instr_low_equal dispath_instr_privilege execute_instr_sub1_low_equal) qed qed next case False then have f_non_exe: "exe_mode_val s1 = False" by auto then have "exe_mode_val s1 = False \ exe_mode_val s2 = False" proof - have "low_equal s1 s2" using a1 by auto then have "state_var s1 = state_var s2" by (simp add: low_equal_def) then have "exe_mode_val s1 = exe_mode_val s2" by (simp add: exe_mode_val_def) then show ?thesis using f_non_exe by auto qed then show ?thesis using f1 a1 apply (simp add: NEXT_def execute_instruction_def) by (simp add: simpler_gets_def bind_def h1_def h2_def Let_def return_def) qed qed -function (sequential) SEQ:: "nat \ ('a::len0) sparc_state \ ('a) sparc_state option" +function (sequential) SEQ:: "nat \ ('a::len) sparc_state \ ('a) sparc_state option" where "SEQ 0 s = Some s" |"SEQ n s = ( case SEQ (n-1) s of None \ None | Some t \ NEXT t )" by pat_completeness auto termination by lexicographic_order lemma SEQ_suc: "SEQ n s = Some t \ SEQ (Suc n) s = NEXT t" apply (induction n) apply clarsimp by (simp add: option.case_eq_if) -definition user_seq_exe:: "nat \ ('a::len0) sparc_state \ bool" where +definition user_seq_exe:: "nat \ ('a::len) sparc_state \ bool" where "user_seq_exe n s \ \i t. (i \ n \ SEQ i s = Some t) \ (good_context t \ get_delayed_pool t = [] \ get_trap_set t = {})" text \NIA is short for non-interference assumption.\ definition "NIA t1 t2 \ ((ucast (get_S (cpu_reg_val PSR t1)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR t2)))::word1) = 0 \ good_context t1 \ get_delayed_pool t1 = [] \ get_trap_set t1 = {} \ good_context t2 \ get_delayed_pool t2 = [] \ get_trap_set t2 = {} \ low_equal t1 t2" text \NIC is short for non-interference conclusion.\ definition "NIC t1 t2 \ (\u1 u2. Some u1 = NEXT t1 \ Some u2 = NEXT t2 \ ((ucast (get_S (cpu_reg_val PSR u1)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR u2)))::word1) = 0 \ low_equal u1 u2)" lemma NIS_short: "\t1 t2. NIA t1 t2 \ NIC t1 t2" apply (simp add: NIA_def NIC_def) using non_interference_step by auto lemma non_interference_induct_case_sub1: assumes a1: "(\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ NIA t1 t2))" shows "(\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ NIA t1 t2 \ NIC t1 t2))" using NIS_short using assms by auto lemma non_interference_induct_case: assumes a1: "((\i t. i \ n \ SEQ i s1 = Some t \ good_context t \ get_delayed_pool t = [] \ get_trap_set t = {}) \ (\i t. i \ n \ SEQ i s2 = Some t \ good_context t \ get_delayed_pool t = [] \ get_trap_set t = {}) \ (\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ ((ucast (get_S (cpu_reg_val PSR t1)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR t2)))::word1) = 0 \ low_equal t1 t2))) \ (\i t. i \ Suc n \ SEQ i s1 = Some t \ good_context t \ get_delayed_pool t = [] \ get_trap_set t = {}) \ (\i t. i \ Suc n \ SEQ i s2 = Some t \ good_context t \ get_delayed_pool t = [] \ get_trap_set t = {})" shows "\t1. Some t1 = (case SEQ n s1 of None \ None | Some x \ NEXT x) \ (\t2. Some t2 = (case SEQ n s2 of None \ None | Some x \ NEXT x) \ ((ucast (get_S (cpu_reg_val PSR t1)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR t2)))::word1) = 0 \ low_equal t1 t2)" proof - from a1 have f1: "((\i t. i \ n \ SEQ i s1 = Some t \ good_context t \ get_delayed_pool t = [] \ get_trap_set t = {}) \ (\i t. i \ n \ SEQ i s2 = Some t \ good_context t \ get_delayed_pool t = [] \ get_trap_set t = {}))" by (metis le_SucI) then have f2: "(\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ ((ucast (get_S (cpu_reg_val PSR t1)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR t2)))::word1) = 0 \ low_equal t1 t2))" using a1 by auto then have f3: "(\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ NIA t1 t2))" using f1 NIA_def by (metis (full_types) dual_order.refl) then have "(\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ NIA t1 t2 \ NIC t1 t2))" using non_interference_induct_case_sub1 by blast then have "(\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ (((ucast (get_S (cpu_reg_val PSR t1)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR t2)))::word1) = 0 \ good_context t1 \ get_delayed_pool t1 = [] \ get_trap_set t1 = {} \ good_context t2 \ get_delayed_pool t2 = [] \ get_trap_set t2 = {} \ low_equal t1 t2) \ (\u1 u2. Some u1 = NEXT t1 \ Some u2 = NEXT t2 \ ((ucast (get_S (cpu_reg_val PSR u1)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR u2)))::word1) = 0 \ low_equal u1 u2)))" using NIA_def NIC_def by fastforce then show ?thesis by (metis option.simps(5)) qed lemma non_interference_induct_case_sub2: assumes a1: "(user_seq_exe n s1 \ user_seq_exe n s2 \ (\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ ((ucast (get_S (cpu_reg_val PSR t1)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR t2)))::word1) = 0 \ low_equal t1 t2))) \ user_seq_exe (Suc n) s1 \ user_seq_exe (Suc n) s2" shows "\t1. Some t1 = (case SEQ n s1 of None \ None | Some x \ NEXT x) \ (\t2. Some t2 = (case SEQ n s2 of None \ None | Some x \ NEXT x) \ ((ucast (get_S (cpu_reg_val PSR t1)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR t2)))::word1) = 0 \ low_equal t1 t2)" using a1 by (simp add: non_interference_induct_case user_seq_exe_def) theorem non_interference: assumes a1: "((ucast (get_S (cpu_reg_val PSR s1)))::word1) = 0 \ good_context s1 \ get_delayed_pool s1 = [] \ get_trap_set s1 = {} \ ((ucast (get_S (cpu_reg_val PSR s2)))::word1) = 0 \ get_delayed_pool s2 = [] \ get_trap_set s2 = {} \ good_context s2 \ user_seq_exe n s1 \ user_seq_exe n s2 \ low_equal s1 s2" shows "(\t1 t2. Some t1 = SEQ n s1 \ Some t2 = SEQ n s2 \ ((ucast (get_S (cpu_reg_val PSR t1)))::word1) = 0 \ ((ucast (get_S (cpu_reg_val PSR t2)))::word1) = 0 \ low_equal t1 t2)" using a1 apply (induction n) apply (simp add: user_seq_exe_def) apply clarsimp by (simp add: non_interference_induct_case_sub2) end diff --git a/thys/SPARCv8/SparcModel_MMU/Sparc_State.thy b/thys/SPARCv8/SparcModel_MMU/Sparc_State.thy --- a/thys/SPARCv8/SparcModel_MMU/Sparc_State.thy +++ b/thys/SPARCv8/SparcModel_MMU/Sparc_State.thy @@ -1,1007 +1,1007 @@ (* * Copyright 2016, NTU * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * Author: Zhe Hou, David Sanan. *) section \SPARC V8 state model\ theory Sparc_State imports Main Sparc_Types "../lib/wp/DetMonadLemmas" MMU begin section \state as a function\ record cpu_cache = dcache:: cache_context icache:: cache_context text\ The state @{term sparc_state} is defined as a tuple @{term cpu_context}, @{term user_context}, @{term mem_context}, defining the state of the CPU registers, user registers, memory, cache, and delayed write pool respectively. Additionally, a boolean indicates whether the state is undefined or not. \ record (overloaded) ('a) sparc_state = cpu_reg:: cpu_context user_reg:: "('a) user_context" sys_reg:: sys_context mem:: mem_context mmu:: MMU_state cache:: cpu_cache dwrite:: delayed_write_pool state_var:: sparc_state_var traps:: "Trap set" undef:: bool section\functions for state member access\ definition cpu_reg_val:: "CPU_register \ ('a) sparc_state \ reg_type" where "cpu_reg_val reg state \ (cpu_reg state) reg" definition cpu_reg_mod :: "word32 \ CPU_register \ ('a) sparc_state \ ('a) sparc_state" where "cpu_reg_mod data_w32 cpu state \ state\cpu_reg := ((cpu_reg state)(cpu := data_w32))\" text \r[0] = 0. Otherwise read the actual value.\ definition user_reg_val:: "('a) window_size \ user_reg_type \ ('a) sparc_state \ reg_type" where "user_reg_val window ur state \ if ur = 0 then 0 else (user_reg state) window ur" text \Write a global register. win should be initialised as NWINDOWS.\ fun (sequential) global_reg_mod :: "word32 \ nat \ user_reg_type \ - ('a::len0) sparc_state \ ('a) sparc_state" + ('a::len) sparc_state \ ('a) sparc_state" where "global_reg_mod data_w32 0 ur state = state" | "global_reg_mod data_w32 win ur state = ( let win_word = word_of_int (int (win-1)); ns = state\user_reg := (user_reg state)(win_word := ((user_reg state) win_word)(ur := data_w32))\ in global_reg_mod data_w32 (win-1) ur ns )" text \Compute the next window.\ -definition next_window :: "('a::len0) window_size \ ('a) window_size" +definition next_window :: "('a::len) window_size \ ('a) window_size" where "next_window win \ if (uint win) < (NWINDOWS - 1) then (win + 1) else 0 " text \Compute the previous window.\ -definition pre_window :: "('a::len0) window_size \ ('a::len0) window_size" +definition pre_window :: "('a::len) window_size \ ('a::len) window_size" where "pre_window win \ if (uint win) > 0 then (win - 1) else (word_of_int (NWINDOWS - 1)) " text \write an output register. Also write ur+16 of the previous window.\ -definition out_reg_mod :: "word32 \ ('a::len0) window_size \ user_reg_type \ +definition out_reg_mod :: "word32 \ ('a::len) window_size \ user_reg_type \ ('a) sparc_state \ ('a) sparc_state" where "out_reg_mod data_w32 win ur state \ let state' = state\user_reg := (user_reg state)(win := ((user_reg state) win)(ur := data_w32))\; win' = pre_window win; ur' = ur + 16 in state'\user_reg := (user_reg state')(win' := ((user_reg state') win')(ur' := data_w32))\ " text \Write a input register. Also write ur-16 of the next window.\ -definition in_reg_mod :: "word32 \ ('a::len0) window_size \ user_reg_type \ +definition in_reg_mod :: "word32 \ ('a::len) window_size \ user_reg_type \ ('a) sparc_state \ ('a) sparc_state" where "in_reg_mod data_w32 win ur state \ let state' = state\user_reg := (user_reg state)(win := ((user_reg state) win)(ur := data_w32))\; win' = next_window win; ur' = ur - 16 in state'\user_reg := (user_reg state')(win' := ((user_reg state') win')(ur' := data_w32))\ " text \Do not modify r[0].\ -definition user_reg_mod :: "word32 \ ('a::len0) window_size \ user_reg_type \ +definition user_reg_mod :: "word32 \ ('a::len) window_size \ user_reg_type \ ('a) sparc_state \ ('a) sparc_state" where "user_reg_mod data_w32 win ur state \ if ur = 0 then state else if 0 < ur \ ur < 8 then global_reg_mod data_w32 (nat NWINDOWS) ur state else if 7 < ur \ ur < 16 then out_reg_mod data_w32 win ur state else if 15 < ur \ ur < 24 then state\user_reg := (user_reg state)(win := ((user_reg state) win)(ur := data_w32))\ else \<^cancel>\if 23 < ur \ ur < 32 then\ in_reg_mod data_w32 win ur state \<^cancel>\else state\ " definition sys_reg_val :: "sys_reg \ ('a) sparc_state \ reg_type" where "sys_reg_val reg state \ (sys_reg state) reg" definition sys_reg_mod :: "word32 \ sys_reg \ ('a) sparc_state \ ('a) sparc_state" where "sys_reg_mod data_w32 sys state \ state\sys_reg := (sys_reg state)(sys := data_w32)\" text \The following fucntions deal with physical memory. N.B. Physical memory address in SPARCv8 is 36-bit.\ text \LEON3 doesn't distinguish ASI 8 and 9; 10 and 11 for read access for both user and supervisor. We recently discovered that the compiled machine code by the sparc-elf compiler often reads asi = 10 (user data) when the actual content is store in asi = 8 (user instruction). For testing purposes, we don't distinguish asi = 8,9,10,11 for reading access.\ definition mem_val:: "asi_type \ phys_address \ ('a) sparc_state \ mem_val_type option" where "mem_val asi add state \ let asi8 = word_of_int 8; asi9 = word_of_int 9; asi10 = word_of_int 10; asi11 = word_of_int 11; r1 = (mem state) asi8 add in if r1 = None then let r2 = (mem state) asi9 add in if r2 = None then let r3 = (mem state) asi10 add in if r3 = None then (mem state) asi11 add else r3 else r2 else r1 " text \An alternative way to read values from memory. Some implementations may use this definition.\ definition mem_val_alt:: "asi_type \ phys_address \ ('a) sparc_state \ mem_val_type option" where "mem_val_alt asi add state \ let r1 = (mem state) asi add; asi8 = word_of_int 8; asi9 = word_of_int 9; asi10 = word_of_int 10; asi11 = word_of_int 11 in if r1 = None \ (uint asi) = 8 then let r2 = (mem state) asi9 add in r2 else if r1 = None \ (uint asi) = 9 then let r2 = (mem state) asi8 add in r2 else if r1 = None \ (uint asi) = 10 then let r2 = (mem state) asi11 add in if r2 = None then let r3 = (mem state) asi8 add in if r3 = None then (mem state) asi9 add else r3 else r2 else if r1 = None \ (uint asi) = 11 then let r2 = (mem state) asi10 add in if r2 = None then let r3 = (mem state) asi8 add in if r3 = None then (mem state) asi9 add else r3 else r2 else r1" definition mem_mod :: "asi_type \ phys_address \ mem_val_type \ ('a) sparc_state \ ('a) sparc_state" where "mem_mod asi addr val state \ let state1 = state\mem := (mem state) (asi := ((mem state) asi)(addr := Some val))\ in \ \Only allow one of \asi\ 8 and 9 (10 and 11) to have value.\ if (uint asi) = 8 \ (uint asi) = 10 then let asi2 = word_of_int ((uint asi) + 1) in state1\mem := (mem state1) (asi2 := ((mem state1) asi2)(addr := None))\ else if (uint asi) = 9 \ (uint asi) = 11 then let asi2 = word_of_int ((uint asi) - 1) in state1\mem := (mem state1)(asi2 := ((mem state1) asi2)(addr := None))\ else state1 " text \An alternative way to write memory. This method insists that for each address, it can only hold a value in one of ASI = 8,9,10,11.\ definition mem_mod_alt :: "asi_type \ phys_address \ mem_val_type \ ('a) sparc_state \ ('a) sparc_state" where "mem_mod_alt asi addr val state \ let state1 = state\mem := (mem state) (asi := ((mem state) asi)(addr := Some val))\; asi8 = word_of_int 8; asi9 = word_of_int 9; asi10 = word_of_int 10; asi11 = word_of_int 11 in \ \Only allow one of \asi\ 8, 9, 10, 11 to have value.\ if (uint asi) = 8 then let state2 = state1\mem := (mem state1) (asi9 := ((mem state1) asi9)(addr := None))\; state3 = state2\mem := (mem state2) (asi10 := ((mem state2) asi10)(addr := None))\; state4 = state3\mem := (mem state3) (asi11 := ((mem state3) asi11)(addr := None))\ in state4 else if (uint asi) = 9 then let state2 = state1\mem := (mem state1) (asi8 := ((mem state1) asi8)(addr := None))\; state3 = state2\mem := (mem state2) (asi10 := ((mem state2) asi10)(addr := None))\; state4 = state3\mem := (mem state3) (asi11 := ((mem state3) asi11)(addr := None))\ in state4 else if (uint asi) = 10 then let state2 = state1\mem := (mem state1) (asi9 := ((mem state1) asi9)(addr := None))\; state3 = state2\mem := (mem state2) (asi8 := ((mem state2) asi8)(addr := None))\; state4 = state3\mem := (mem state3) (asi11 := ((mem state3) asi11)(addr := None))\ in state4 else if (uint asi) = 11 then let state2 = state1\mem := (mem state1) (asi9 := ((mem state1) asi9)(addr := None))\; state3 = state2\mem := (mem state2) (asi10 := ((mem state2) asi10)(addr := None))\; state4 = state3\mem := (mem state3) (asi8 := ((mem state3) asi8)(addr := None))\ in state4 else state1 " text \Given an ASI (word8), an address (word32) addr, read the 32bit value from the memory addresses starting from address addr' where addr' = addr exception that the last two bits are 0's. That is, read the data from addr', addr'+1, addr'+2, addr'+3.\ definition mem_val_w32 :: "asi_type \ phys_address \ ('a) sparc_state \ word32 option" where "mem_val_w32 asi addr state \ let addr' = (AND) addr 0b111111111111111111111111111111111100; addr0 = addr'; addr1 = addr' + 1; addr2 = addr' + 2; addr3 = addr' + 3; r0 = mem_val_alt asi addr0 state; r1 = mem_val_alt asi addr1 state; r2 = mem_val_alt asi addr2 state; r3 = mem_val_alt asi addr3 state in if r0 = None \ r1 = None \ r2 = None \ r3 = None then None else let byte0 = case r0 of Some v \ v; byte1 = case r1 of Some v \ v; byte2 = case r2 of Some v \ v; byte3 = case r3 of Some v \ v in Some ((OR) ((OR) ((OR) ((ucast(byte0)) << 24) ((ucast(byte1)) << 16)) ((ucast(byte2)) << 8)) (ucast(byte3))) " text \ Let \addr'\ be \addr\ with last two bits set to 0's. Write the 32bit data in the memory address \addr'\ (and the following 3 addresses). \byte_mask\ decides which byte of the 32bits are written. \ definition mem_mod_w32 :: "asi_type \ phys_address \ word4 \ word32 \ ('a) sparc_state \ ('a) sparc_state" where "mem_mod_w32 asi addr byte_mask data_w32 state \ let addr' = (AND) addr 0b111111111111111111111111111111111100; addr0 = (OR) addr' 0b000000000000000000000000000000000000; addr1 = (OR) addr' 0b000000000000000000000000000000000001; addr2 = (OR) addr' 0b000000000000000000000000000000000010; addr3 = (OR) addr' 0b000000000000000000000000000000000011; byte0 = (ucast (data_w32 >> 24))::mem_val_type; byte1 = (ucast (data_w32 >> 16))::mem_val_type; byte2 = (ucast (data_w32 >> 8))::mem_val_type; byte3 = (ucast data_w32)::mem_val_type; s0 = if (((AND) byte_mask (0b1000::word4)) >> 3) = 1 then mem_mod asi addr0 byte0 state else state; s1 = if (((AND) byte_mask (0b0100::word4)) >> 2) = 1 then mem_mod asi addr1 byte1 s0 else s0; s2 = if (((AND) byte_mask (0b0010::word4)) >> 1) = 1 then mem_mod asi addr2 byte2 s1 else s1; s3 = if ((AND) byte_mask (0b0001::word4)) = 1 then mem_mod asi addr3 byte3 s2 else s2 in s3 " text \The following functions deal with virtual addresses. These are based on functions written by David Sanan.\ definition load_word_mem :: "('a) sparc_state \ virtua_address \ asi_type \ machine_word option" where "load_word_mem state va asi \ let pair = (virt_to_phys va (mmu state) (mem state)) in case pair of Some pair \ ( if mmu_readable (get_acc_flag (snd pair)) asi then (mem_val_w32 asi (fst pair) state) else None) | None \ None" definition store_word_mem ::"('a) sparc_state \ virtua_address \ machine_word \ word4 \ asi_type \ ('a) sparc_state option" where "store_word_mem state va wd byte_mask asi \ let pair = (virt_to_phys va (mmu state) (mem state)) in case pair of Some pair \ ( if mmu_writable (get_acc_flag (snd pair)) asi then Some (mem_mod_w32 asi (fst pair) byte_mask wd state) else None) | None \ None" definition icache_val:: "cache_type \ ('a) sparc_state \ mem_val_type option" where "icache_val c state \ icache (cache state) c" definition dcache_val:: "cache_type \ ('a) sparc_state \ mem_val_type option" where "dcache_val c state \ dcache (cache state) c" definition icache_mod :: "cache_type \ mem_val_type \ ('a) sparc_state \ ('a) sparc_state" where "icache_mod c val state \ state\cache := ((cache state) \icache := (icache (cache state))(c := Some val)\)\ " definition dcache_mod :: "cache_type \ mem_val_type \ ('a) sparc_state \ ('a) sparc_state" where "dcache_mod c val state \ state\cache := ((cache state) \dcache := (dcache (cache state))(c := Some val)\)\ " text \Check if the memory address is in the cache or not.\ definition icache_miss :: "virtua_address \ ('a) sparc_state \ bool" where "icache_miss addr state \ let line_len = 12; tag = (ucast (addr >> line_len))::cache_tag; line = (ucast (0b0::word1))::cache_line_size in if (icache_val (tag,line) state) = None then True else False " text \Check if the memory address is in the cache or not.\ definition dcache_miss :: "virtua_address \ ('a) sparc_state \ bool" where "dcache_miss addr state \ let line_len = 12; tag = (ucast (addr >> line_len))::cache_tag; line = (ucast (0b0::word1))::cache_line_size in if (dcache_val (tag,line) state) = None then True else False " definition read_data_cache:: "('a) sparc_state \ virtua_address \ machine_word option" where "read_data_cache state va \ let tag = (ucast (va >> 12))::word20; offset0 = (AND) ((ucast va)::word12) 0b111111111100; offset1 = (OR) offset0 0b000000000001; offset2 = (OR) offset0 0b000000000010; offset3 = (OR) offset0 0b000000000011; r0 = dcache_val (tag,offset0) state; r1 = dcache_val (tag,offset1) state; r2 = dcache_val (tag,offset2) state; r3 = dcache_val (tag,offset3) state in if r0 = None \ r1 = None \ r2 = None \ r3 = None then None else let byte0 = case r0 of Some v \ v; byte1 = case r1 of Some v \ v; byte2 = case r2 of Some v \ v; byte3 = case r3 of Some v \ v in Some ((OR) ((OR) ((OR) ((ucast(byte0)) << 24) ((ucast(byte1)) << 16)) ((ucast(byte2)) << 8)) (ucast(byte3))) " definition read_instr_cache:: "('a) sparc_state \ virtua_address \ machine_word option" where "read_instr_cache state va \ let tag = (ucast (va >> 12))::word20; offset0 = (AND) ((ucast va)::word12) 0b111111111100; offset1 = (OR) offset0 0b000000000001; offset2 = (OR) offset0 0b000000000010; offset3 = (OR) offset0 0b000000000011; r0 = icache_val (tag,offset0) state; r1 = icache_val (tag,offset1) state; r2 = icache_val (tag,offset2) state; r3 = icache_val (tag,offset3) state in if r0 = None \ r1 = None \ r2 = None \ r3 = None then None else let byte0 = case r0 of Some v \ v; byte1 = case r1 of Some v \ v; byte2 = case r2 of Some v \ v; byte3 = case r3 of Some v \ v in Some ((OR) ((OR) ((OR) ((ucast(byte0)) << 24) ((ucast(byte1)) << 16)) ((ucast(byte2)) << 8)) (ucast(byte3))) " definition add_data_cache :: "('a) sparc_state \ virtua_address \ machine_word \ word4 \ ('a) sparc_state" where "add_data_cache state va word byte_mask \ let tag = (ucast (va >> 12))::word20; offset0 = (AND) ((ucast va)::word12) 0b111111111100; offset1 = (OR) offset0 0b000000000001; offset2 = (OR) offset0 0b000000000010; offset3 = (OR) offset0 0b000000000011; byte0 = (ucast (word >> 24))::mem_val_type; byte1 = (ucast (word >> 16))::mem_val_type; byte2 = (ucast (word >> 8))::mem_val_type; byte3 = (ucast word)::mem_val_type; s0 = if (((AND) byte_mask (0b1000::word4)) >> 3) = 1 then dcache_mod (tag,offset0) byte0 state else state; s1 = if (((AND) byte_mask (0b0100::word4)) >> 2) = 1 then dcache_mod (tag,offset1) byte1 s0 else s0; s2 = if (((AND) byte_mask (0b0010::word4)) >> 1) = 1 then dcache_mod (tag,offset2) byte2 s1 else s1; s3 = if ((AND) byte_mask (0b0001::word4)) = 1 then dcache_mod (tag,offset3) byte3 s2 else s2 in s3 " definition add_instr_cache :: "('a) sparc_state \ virtua_address \ machine_word \ word4 \ ('a) sparc_state" where "add_instr_cache state va word byte_mask \ let tag = (ucast (va >> 12))::word20; offset0 = (AND) ((ucast va)::word12) 0b111111111100; offset1 = (OR) offset0 0b000000000001; offset2 = (OR) offset0 0b000000000010; offset3 = (OR) offset0 0b000000000011; byte0 = (ucast (word >> 24))::mem_val_type; byte1 = (ucast (word >> 16))::mem_val_type; byte2 = (ucast (word >> 8))::mem_val_type; byte3 = (ucast word)::mem_val_type; s0 = if (((AND) byte_mask (0b1000::word4)) >> 3) = 1 then icache_mod (tag,offset0) byte0 state else state; s1 = if (((AND) byte_mask (0b0100::word4)) >> 2) = 1 then icache_mod (tag,offset1) byte1 s0 else s0; s2 = if (((AND) byte_mask (0b0010::word4)) >> 1) = 1 then icache_mod (tag,offset2) byte2 s1 else s1; s3 = if ((AND) byte_mask (0b0001::word4)) = 1 then icache_mod (tag,offset3) byte3 s2 else s2 in s3 " definition empty_cache ::"cache_context" where "empty_cache c \ None" definition flush_data_cache:: "('a) sparc_state \ ('a) sparc_state" where "flush_data_cache state \ state\cache := ((cache state)\dcache := empty_cache\)\" definition flush_instr_cache:: "('a) sparc_state \ ('a) sparc_state" where "flush_instr_cache state \ state\cache := ((cache state)\icache := empty_cache\)\" definition flush_cache_all:: "('a) sparc_state \ ('a) sparc_state" where "flush_cache_all state \ state\cache := ((cache state)\ icache := empty_cache, dcache := empty_cache\)\" text \Check if the FI or FD bit of CCR is 1. If FI is 1 then flush instruction cache. If FD is 1 then flush data cache.\ definition ccr_flush :: "('a) sparc_state \ ('a) sparc_state" where "ccr_flush state \ let ccr_val = sys_reg_val CCR state; \ \\FI\ is bit 21 of \CCR\\ fi_val = ((AND) ccr_val (0b00000000001000000000000000000000)) >> 21; fd_val = ((AND) ccr_val (0b00000000010000000000000000000000)) >> 22; state1 = (if fi_val = 1 then flush_instr_cache state else state) in if fd_val = 1 then flush_data_cache state1 else state1" definition get_delayed_pool :: "('a) sparc_state \ delayed_write_pool" where "get_delayed_pool state \ dwrite state" definition exe_pool :: "(int \ reg_type \ CPU_register) \ (int \ reg_type \ CPU_register)" where "exe_pool w \ case w of (n,v,c) \ ((n-1),v,c)" text \Minus 1 to the delayed count for all the members in the set. Assuming all members have delay > 0.\ primrec delayed_pool_minus :: "delayed_write_pool \ delayed_write_pool" where "delayed_pool_minus [] = []" | "delayed_pool_minus (x#xs) = (exe_pool x)#(delayed_pool_minus xs)" text \Add a delayed-write to the pool.\ definition delayed_pool_add :: "(int \ reg_type \ CPU_register) \ ('a) sparc_state \ ('a) sparc_state" where "delayed_pool_add dw s \ let (i,v,cr) = dw in if i = 0 then \ \Write the value to the register immediately.\ cpu_reg_mod v cr s else \ \Add to delayed write pool.\ let curr_pool = get_delayed_pool s in s\dwrite := curr_pool@[dw]\" text \Remove a delayed-write from the pool. Assume that the delayed-write to be removed has delay 0. i.e., it has been executed.\ definition delayed_pool_rm :: "(int \ reg_type \ CPU_register) \ ('a) sparc_state \ ('a) sparc_state" where "delayed_pool_rm dw s \ let curr_pool = get_delayed_pool s in case dw of (n,v,cr) \ (if n = 0 then s\dwrite := List.remove1 dw curr_pool\ else s) " text \Remove all the entries with delay = 0, i.e., those that are written.\ primrec delayed_pool_rm_written :: "delayed_write_pool \ delayed_write_pool" where "delayed_pool_rm_written [] = []" | "delayed_pool_rm_written (x#xs) = (if fst x = 0 then delayed_pool_rm_written xs else x#(delayed_pool_rm_written xs)) " definition annul_val :: "('a) sparc_state \ bool" where "annul_val state \ get_annul (state_var state)" definition annul_mod :: "bool \ ('a) sparc_state \ ('a) sparc_state" where "annul_mod b s \ s\state_var := write_annul b (state_var s)\" definition reset_trap_val :: "('a) sparc_state \ bool" where "reset_trap_val state \ get_reset_trap (state_var state)" definition reset_trap_mod :: "bool \ ('a) sparc_state \ ('a) sparc_state" where "reset_trap_mod b s \ s\state_var := write_reset_trap b (state_var s)\" definition exe_mode_val :: "('a) sparc_state \ bool" where "exe_mode_val state \ get_exe_mode (state_var state)" definition exe_mode_mod :: "bool \ ('a) sparc_state \ ('a) sparc_state" where "exe_mode_mod b s \ s\state_var := write_exe_mode b (state_var s)\" definition reset_mode_val :: "('a) sparc_state \ bool" where "reset_mode_val state \ get_reset_mode (state_var state)" definition reset_mode_mod :: "bool \ ('a) sparc_state \ ('a) sparc_state" where "reset_mode_mod b s \ s\state_var := write_reset_mode b (state_var s)\" definition err_mode_val :: "('a) sparc_state \ bool" where "err_mode_val state \ get_err_mode (state_var state)" definition err_mode_mod :: "bool \ ('a) sparc_state \ ('a) sparc_state" where "err_mode_mod b s \ s\state_var := write_err_mode b (state_var s)\" definition ticc_trap_type_val :: "('a) sparc_state \ word7" where "ticc_trap_type_val state \ get_ticc_trap_type (state_var state)" definition ticc_trap_type_mod :: "word7 \ ('a) sparc_state \ ('a) sparc_state" where "ticc_trap_type_mod w s \ s\state_var := write_ticc_trap_type w (state_var s)\" definition interrupt_level_val :: "('a) sparc_state \ word3" where "interrupt_level_val state \ get_interrupt_level (state_var state)" definition interrupt_level_mod :: "word3 \ ('a) sparc_state \ ('a) sparc_state" where "interrupt_level_mod w s \ s\state_var := write_interrupt_level w (state_var s)\" definition store_barrier_pending_val :: "('a) sparc_state \ bool" where "store_barrier_pending_val state \ get_store_barrier_pending (state_var state)" definition store_barrier_pending_mod :: "bool \ ('a) sparc_state \ ('a) sparc_state" where "store_barrier_pending_mod w s \ s\state_var := write_store_barrier_pending w (state_var s)\" definition pb_block_ldst_byte_val :: "virtua_address \ ('a) sparc_state \ bool" where "pb_block_ldst_byte_val add state \ (atm_ldst_byte (state_var state)) add" definition pb_block_ldst_byte_mod :: "virtua_address \ bool \ ('a) sparc_state \ ('a) sparc_state" where "pb_block_ldst_byte_mod add b s \ s\state_var := ((state_var s) \atm_ldst_byte := (atm_ldst_byte (state_var s))(add := b)\)\" text \We only read the address such that add mod 4 = 0. add mod 4 represents the current word.\ definition pb_block_ldst_word_val :: "virtua_address \ ('a) sparc_state \ bool" where "pb_block_ldst_word_val add state \ let add0 = ((AND) add (0b11111111111111111111111111111100::word32)) in (atm_ldst_word (state_var state)) add0" text \We only write the address such that add mod 4 = 0. add mod 4 represents the current word.\ definition pb_block_ldst_word_mod :: "virtua_address \ bool \ ('a) sparc_state \ ('a) sparc_state" where "pb_block_ldst_word_mod add b s \ let add0 = ((AND) add (0b11111111111111111111111111111100::word32)) in s\state_var := ((state_var s) \atm_ldst_word := (atm_ldst_word (state_var s))(add0 := b)\)\" definition get_trap_set :: "('a) sparc_state \ Trap set" where "get_trap_set state \ (traps state)" definition add_trap_set :: "Trap \ ('a) sparc_state \ ('a) sparc_state" where "add_trap_set t s \ s\traps := (traps s) \ {t}\" definition emp_trap_set :: "('a) sparc_state \ ('a) sparc_state" where "emp_trap_set s \ s\traps := {}\" definition state_undef:: "('a) sparc_state \ bool" where "state_undef state \ (undef state)" text \The \memory_read\ interface that conforms with the SPARCv8 manual.\ definition memory_read :: "asi_type \ virtua_address \ ('a) sparc_state \ ((word32 option) \ ('a) sparc_state)" where "memory_read asi addr state \ let asi_int = uint asi in \ \See Page 25 and 35 for ASI usage in LEON 3FT.\ if asi_int = 1 then \ \Forced cache miss.\ \ \Directly read from memory.\ let r1 = load_word_mem state addr (word_of_int 8) in if r1 = None then let r2 = load_word_mem state addr (word_of_int 10) in if r2 = None then (None,state) else (r2,state) else (r1,state) else if asi_int = 2 then \ \System registers.\ \ \See Table 19, Page 34 for System Register address map in LEON 3FT.\ if uint addr = 0 then \ \Cache control register.\ ((Some (sys_reg_val CCR state)), state) else if uint addr = 8 then \ \Instruction cache configuration register.\ ((Some (sys_reg_val ICCR state)), state) else if uint addr = 12 then \ \Data cache configuration register.\ ((Some (sys_reg_val DCCR state)), state) else \ \Invalid address.\ (None, state) else if asi_int \ {8,9} then \ \Access instruction memory.\ let ccr_val = (sys_reg state) CCR in if ccr_val AND 1 \ 0 then \ \Cache is enabled. Update cache.\ \ \We don't go through the tradition, i.e., read from cache first,\ \ \if the address is not cached, then read from memory,\ \ \because performance is not an issue here.\ \ \Thus we directly read from memory and update the cache.\ let data = load_word_mem state addr asi in case data of Some w \ (Some w,(add_instr_cache state addr w (0b1111::word4))) |None \ (None, state) else \ \Cache is disabled. Just read from memory.\ ((load_word_mem state addr asi),state) else if asi_int \ {10,11} then \ \Access data memory.\ let ccr_val = (sys_reg state) CCR in if ccr_val AND 1 \ 0 then \ \Cache is enabled. Update cache.\ \ \We don't go through the tradition, i.e., read from cache first,\ \ \if the address is not cached, then read from memory,\ \ \because performance is not an issue here.\ \ \Thus we directly read from memory and update the cache.\ let data = load_word_mem state addr asi in case data of Some w \ (Some w,(add_data_cache state addr w (0b1111::word4))) |None \ (None, state) else \ \Cache is disabled. Just read from memory.\ ((load_word_mem state addr asi),state) \ \We don't access instruction cache tag. i.e., \asi = 12\.\ else if asi_int = 13 then \ \Read instruction cache data.\ let cache_result = read_instr_cache state addr in case cache_result of Some w \ (Some w, state) |None \ (None, state) \ \We don't access data cache tag. i.e., \asi = 14\.\ else if asi_int = 15 then \ \Read data cache data.\ let cache_result = read_data_cache state addr in case cache_result of Some w \ (Some w, state) |None \ (None, state) else if asi_int \ {16,17} then \ \Flush entire instruction/data cache.\ (None, state) \ \Has no effect for memory read.\ else if asi_int \ {20,21} then \ \MMU diagnostic cache access.\ (None, state) \ \Not considered in this model.\ else if asi_int = 24 then \ \Flush cache and TLB in LEON3.\ \ \But is not used for memory read.\ (None, state) else if asi_int = 25 then \ \MMU registers.\ \ \Treat MMU registers as memory addresses that are not in the main memory.\ ((mmu_reg_val (mmu state) addr), state) else if asi_int = 28 then \ \MMU bypass.\ \ \Directly use addr as a physical address.\ \ \Append 0000 in the front of addr.\ \ \In this case, (ucast addr) suffices.\ ((mem_val_w32 asi (ucast addr) state), state) else if asi_int = 29 then \ \MMU diagnostic access.\ (None, state) \ \Not considered in this model.\ else \ \Not considered in this model.\ (None, state) " text \Get the value of a memory address and an ASI.\ definition mem_val_asi:: "asi_type \ phys_address \ ('a) sparc_state \ mem_val_type option" where "mem_val_asi asi add state \ (mem state) asi add" text \Check if an address is used in ASI 9 or 11.\ definition sup_addr :: "phys_address \ ('a) sparc_state \ bool" where "sup_addr addr state \ let addr' = (AND) addr 0b111111111111111111111111111111111100; addr0 = (OR) addr' 0b000000000000000000000000000000000000; addr1 = (OR) addr' 0b000000000000000000000000000000000001; addr2 = (OR) addr' 0b000000000000000000000000000000000010; addr3 = (OR) addr' 0b000000000000000000000000000000000011; r0 = mem_val_asi 9 addr0 state; r1 = mem_val_asi 9 addr1 state; r2 = mem_val_asi 9 addr2 state; r3 = mem_val_asi 9 addr3 state; r4 = mem_val_asi 11 addr0 state; r5 = mem_val_asi 11 addr1 state; r6 = mem_val_asi 11 addr2 state; r7 = mem_val_asi 11 addr3 state in if r0 = None \ r1 = None \ r2 = None \ r3 = None \ r4 = None \ r5 = None \ r6 = None \ r7 = None then False else True " text \The \memory_write\ interface that conforms with SPARCv8 manual.\ text \LEON3 forbids user to write an address in ASI 9 and 11.\ definition memory_write_asi :: "asi_type \ virtua_address \ word4 \ word32 \ ('a) sparc_state \ ('a) sparc_state option" where "memory_write_asi asi addr byte_mask data_w32 state \ let asi_int = uint asi; \ \See Page 25 and 35 for ASI usage in LEON 3FT.\ psr_val = cpu_reg_val PSR state; s_val = get_S psr_val in if asi_int = 1 then \ \Forced cache miss.\ \ \Directly write to memory.\ \ \Assuming writing into \asi = 10\.\ store_word_mem state addr data_w32 byte_mask (word_of_int 10) else if asi_int = 2 then \ \System registers.\ \ \See Table 19, Page 34 for System Register address map in LEON 3FT.\ if uint addr = 0 then \ \Cache control register.\ let s1 = (sys_reg_mod data_w32 CCR state) in \ \Flush the instruction cache if FI of CCR is 1;\ \ \flush the data cache if FD of CCR is 1.\ Some (ccr_flush s1) else if uint addr = 8 then \ \Instruction cache configuration register.\ Some (sys_reg_mod data_w32 ICCR state) else if uint addr = 12 then \ \Data cache configuration register.\ Some (sys_reg_mod data_w32 DCCR state) else \ \Invalid address.\ None else if asi_int \ {8,9} then \ \Access instruction memory.\ \ \Write to memory. LEON3 does write-through. Both cache and the memory are updated.\ let ns = add_instr_cache state addr data_w32 byte_mask in store_word_mem ns addr data_w32 byte_mask asi else if asi_int \ {10,11} then \ \Access data memory.\ \ \Write to memory. LEON3 does write-through. Both cache and the memory are updated.\ let ns = add_data_cache state addr data_w32 byte_mask in store_word_mem ns addr data_w32 byte_mask asi \ \We don't access instruction cache tag. i.e., \asi = 12\.\ else if asi_int = 13 then \ \Write instruction cache data.\ Some (add_instr_cache state addr data_w32 (0b1111::word4)) \ \We don't access data cache tag. i.e., asi = 14.\ else if asi_int = 15 then \ \Write data cache data.\ Some (add_data_cache state addr data_w32 (0b1111::word4)) else if asi_int = 16 then \ \Flush instruction cache.\ Some (flush_instr_cache state) else if asi_int = 17 then \ \Flush data cache.\ Some (flush_data_cache state) else if asi_int \ {20,21} then \ \MMU diagnostic cache access.\ None \ \Not considered in this model.\ else if asi_int = 24 then \ \Flush TLB and cache in LEON3.\ \ \We don't consider TLB here.\ Some (flush_cache_all state) else if asi_int = 25 then \ \MMU registers.\ \ \Treat MMU registers as memory addresses that are not in the main memory.\ let mmu_state' = mmu_reg_mod (mmu state) addr data_w32 in case mmu_state' of Some mmus \ Some (state\mmu := mmus\) |None \ None else if asi_int = 28 then \ \MMU bypass.\ \ \Write to virtual address as physical address.\ \ \Append 0000 in front of addr.\ Some (mem_mod_w32 asi (ucast addr) byte_mask data_w32 state) else if asi_int = 29 then \ \MMU diagnostic access.\ None \ \Not considered in this model.\ else \ \Not considered in this model.\ None " definition memory_write :: "asi_type \ virtua_address \ word4 \ word32 \ ('a) sparc_state \ ('a) sparc_state option" where "memory_write asi addr byte_mask data_w32 state \ let result = memory_write_asi asi addr byte_mask data_w32 state in case result of None \ None | Some s1 \ Some (store_barrier_pending_mod False s1)" text \monad for sequential operations over the register representation\ type_synonym ('a,'e) sparc_state_monad = "(('a) sparc_state,'e) det_monad" text \Given a word32 value, a cpu register, write the value in the cpu register.\ definition write_cpu :: "word32 \ CPU_register \ ('a,unit) sparc_state_monad" where "write_cpu w cr \ do modify (\s. (cpu_reg_mod w cr s)); return () od" definition write_cpu_tt :: "word8 \ ('a,unit) sparc_state_monad" where "write_cpu_tt w \ do tbr_val \ gets (\s. (cpu_reg_val TBR s)); new_tbr_val \ gets (\s. (write_tt w tbr_val)); write_cpu new_tbr_val TBR; return () od" text \Given a word32 value, a word4 window, a user register, write the value in the user register. N.B. CWP is a 5 bit value, but we only use the last 4 bits, since there are only 16 windows.\ -definition write_reg :: "word32 \ ('a::len0) word \ user_reg_type \ +definition write_reg :: "word32 \ ('a::len) word \ user_reg_type \ ('a,unit) sparc_state_monad" where "write_reg w win ur \ do modify (\s.(user_reg_mod w win ur s)); return () od" definition set_annul :: "bool \ ('a,unit) sparc_state_monad" where "set_annul b \ do modify (\s. (annul_mod b s)); return () od" definition set_reset_trap :: "bool \ ('a,unit) sparc_state_monad" where "set_reset_trap b \ do modify (\s. (reset_trap_mod b s)); return () od" definition set_exe_mode :: "bool \ ('a,unit) sparc_state_monad" where "set_exe_mode b \ do modify (\s. (exe_mode_mod b s)); return () od" definition set_reset_mode :: "bool \ ('a,unit) sparc_state_monad" where "set_reset_mode b \ do modify (\s. (reset_mode_mod b s)); return () od" definition set_err_mode :: "bool \ ('a,unit) sparc_state_monad" where "set_err_mode b \ do modify (\s. (err_mode_mod b s)); return () od" fun get_delayed_0 :: "(int \ reg_type \ CPU_register) list \ (int \ reg_type \ CPU_register) list" where "get_delayed_0 [] = []" | "get_delayed_0 (x # xs) = (if fst x = 0 then x # (get_delayed_0 xs) else get_delayed_0 xs)" text \Get a list of delayed-writes with delay 0.\ definition get_delayed_write :: "delayed_write_pool \ (int \ reg_type \ CPU_register) list" where "get_delayed_write dwp \ get_delayed_0 dwp" definition delayed_write :: "(int \ reg_type \ CPU_register) \ ('a) sparc_state \ ('a) sparc_state" where "delayed_write dw s \ let (n,v,r) = dw in if n = 0 then cpu_reg_mod v r s else s" primrec delayed_write_all :: "(int \ reg_type \ CPU_register) list \ ('a) sparc_state \ ('a) sparc_state" where "delayed_write_all [] s = s" |"delayed_write_all (x # xs) s = delayed_write_all xs (delayed_write x s)" primrec delayed_pool_rm_list :: "(int \ reg_type \ CPU_register) list\ ('a) sparc_state \ ('a) sparc_state" where "delayed_pool_rm_list [] s = s" |"delayed_pool_rm_list (x # xs) s = delayed_pool_rm_list xs (delayed_pool_rm x s)" definition delayed_pool_write :: "('a) sparc_state \ ('a) sparc_state" where "delayed_pool_write s \ let dwp0 = get_delayed_pool s; dwp1 = delayed_pool_minus dwp0; wl = get_delayed_write dwp1; s1 = delayed_write_all wl s; s2 = delayed_pool_rm_list wl s1 in s2" definition raise_trap :: "Trap \ ('a,unit) sparc_state_monad" where "raise_trap t \ do modify (\s. (add_trap_set t s)); return () od" end diff --git a/thys/Word_Lib/Word_EqI.thy b/thys/Word_Lib/Word_EqI.thy --- a/thys/Word_Lib/Word_EqI.thy +++ b/thys/Word_Lib/Word_EqI.thy @@ -1,184 +1,184 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Solving Word Equalities" theory Word_EqI imports Word_Next "HOL-Eisbach.Eisbach_Tools" begin text \ Some word equalities can be solved by considering the problem bitwise for all @{prop "n < LENGTH('a::len)"}, which is different to running @{text word_bitwise} and expanding into an explicit list of bits. \ lemma word_or_zero: "(a || b = 0) = (a = 0 \ b = 0)" by (safe; rule word_eqI, drule_tac x=n in word_eqD, simp) lemma test_bit_over: - "n \ size (x::'a::len0 word) \ (x !! n) = False" + "n \ size (x::'a::len word) \ (x !! n) = False" by (simp add: test_bit_bl word_size) lemma neg_mask_test_bit: "(~~(mask n) :: 'a :: len word) !! m = (n \ m \ m < LENGTH('a))" by (metis not_le nth_mask test_bit_bin word_ops_nth_size word_size) lemma word_2p_mult_inc: assumes x: "2 * 2 ^ n < (2::'a::len word) * 2 ^ m" assumes suc_n: "Suc n < LENGTH('a::len)" shows "2^n < (2::'a::len word)^m" by (smt suc_n le_less_trans lessI nat_less_le nat_mult_less_cancel_disj p2_gt_0 power_Suc power_Suc unat_power_lower word_less_nat_alt x) lemma word_power_increasing: assumes x: "2 ^ x < (2 ^ y::'a::len word)" "x < LENGTH('a::len)" "y < LENGTH('a::len)" shows "x < y" using x apply (induct x arbitrary: y) apply (case_tac y; simp) apply (case_tac y; clarsimp simp: word_2p_mult_inc) apply (subst (asm) power_Suc [symmetric]) apply (subst (asm) p2_eq_0) apply simp done lemma upper_bits_unset_is_l2p: "n < LENGTH('a) \ (\n' \ n. n' < LENGTH('a) \ \ p !! n') = (p < 2 ^ n)" for p :: "'a :: len word" apply (cases "Suc 0 < LENGTH('a)") prefer 2 apply (subgoal_tac "LENGTH('a) = 1", auto simp: word_eq_iff)[1] apply (rule iffI) apply (subst mask_eq_iff_w2p [symmetric]) apply (clarsimp simp: word_size) apply (rule word_eqI, rename_tac n') apply (case_tac "n' < n"; simp add: word_size) by (meson bang_is_le le_less_trans not_le word_power_increasing) lemma less_2p_is_upper_bits_unset: "p < 2 ^ n \ n < LENGTH('a) \ (\n' \ n. n' < LENGTH('a) \ \ p !! n')" for p :: "'a :: len word" by (meson le_less_trans le_mask_iff_lt_2n upper_bits_unset_is_l2p word_zero_le) lemma word_le_minus_one_leq: "x < y \ x \ y - 1" for x :: "'a :: len word" by (simp add: plus_one_helper) lemma word_less_sub_le[simp]: fixes x :: "'a :: len word" assumes nv: "n < LENGTH('a)" shows "(x \ 2 ^ n - 1) = (x < 2 ^ n)" using le_less_trans word_le_minus_one_leq nv power_2_ge_iff by blast lemma not_greatest_aligned: "\ x < y; is_aligned x n; is_aligned y n \ \ x + 2 ^ n \ 0" by (metis NOT_mask add_diff_cancel_right' diff_0 is_aligned_neg_mask_eq not_le word_and_le1) lemma neg_mask_mono_le: "x \ y \ x && ~~(mask n) \ y && ~~(mask n)" for x :: "'a :: len word" proof (rule ccontr, simp add: linorder_not_le, cases "n < LENGTH('a)") case False then show "y && ~~(mask n) < x && ~~(mask n) \ False" by (simp add: mask_def linorder_not_less power_overflow) next case True assume a: "x \ y" and b: "y && ~~(mask n) < x && ~~(mask n)" have word_bits: "n < LENGTH('a)" by fact have "y \ (y && ~~(mask n)) + (y && mask n)" by (simp add: word_plus_and_or_coroll2 add.commute) also have "\ \ (y && ~~(mask n)) + 2 ^ n" apply (rule word_plus_mono_right) apply (rule order_less_imp_le, rule and_mask_less_size) apply (simp add: word_size word_bits) apply (rule is_aligned_no_overflow'', simp add: is_aligned_neg_mask word_bits) apply (rule not_greatest_aligned, rule b; simp add: is_aligned_neg_mask) done also have "\ \ x && ~~(mask n)" using b apply (subst add.commute) apply (rule le_plus) apply (rule aligned_at_least_t2n_diff; simp add: is_aligned_neg_mask) apply (rule ccontr, simp add: linorder_not_le) apply (drule aligned_small_is_0[rotated]; simp add: is_aligned_neg_mask) done also have "\ \ x" by (rule word_and_le2) also have "x \ y" by fact finally show "False" using b by simp qed lemma and_neg_mask_eq_iff_not_mask_le: "w && ~~(mask n) = ~~(mask n) \ ~~(mask n) \ w" by (metis eq_iff neg_mask_mono_le word_and_le1 word_and_le2 word_bw_same(1)) lemma le_mask_high_bits: "w \ mask n \ (\i \ {n ..< size w}. \ w !! i)" by (auto simp: word_size and_mask_eq_iff_le_mask[symmetric] word_eq_iff) lemma neg_mask_le_high_bits: "~~(mask n) \ w \ (\i \ {n ..< size w}. w !! i)" by (auto simp: word_size and_neg_mask_eq_iff_not_mask_le[symmetric] word_eq_iff neg_mask_test_bit) lemma test_bit_conj_lt: "(x !! m \ m < LENGTH('a)) = x !! m" for x :: "'a :: len word" using test_bit_bin by blast lemma neg_test_bit: "(~~ x) !! n = (\ x !! n \ n < LENGTH('a))" for x :: "'a::len word" by (cases "n < LENGTH('a)") (auto simp add: test_bit_over word_ops_nth_size word_size) named_theorems word_eqI_simps lemmas [word_eqI_simps] = word_ops_nth_size word_size word_or_zero neg_mask_test_bit nth_ucast is_aligned_nth nth_w2p nth_shiftl nth_shiftr less_2p_is_upper_bits_unset le_mask_high_bits neg_mask_le_high_bits bang_eq neg_test_bit is_up is_down lemmas word_eqI_rule = word_eqI[rule_format] lemma test_bit_lenD: "x !! n \ n < LENGTH('a) \ x !! n" for x :: "'a :: len word" by (fastforce dest: test_bit_size simp: word_size) method word_eqI uses simp simp_del split split_del cong flip = ((* reduce conclusion to test_bit: *) rule word_eqI_rule, (* make sure we're in clarsimp normal form: *) (clarsimp simp: simp simp del: simp_del simp flip: flip split: split split del: split_del cong: cong)?, (* turn x < 2^n assumptions into mask equations: *) ((drule less_mask_eq)+)?, (* expand and distribute test_bit everywhere: *) (clarsimp simp: word_eqI_simps simp simp del: simp_del simp flip: flip split: split split del: split_del cong: cong)?, (* add any additional word size constraints to new indices: *) ((drule test_bit_lenD)+)?, (* try to make progress (can't use +, would loop): *) (clarsimp simp: word_eqI_simps simp simp del: simp_del simp flip: flip split: split split del: split_del cong: cong)?, (* helps sometimes, rarely: *) (simp add: simp test_bit_conj_lt del: simp_del flip: flip split: split split del: split_del cong: cong)?) method word_eqI_solve uses simp simp_del split split_del cong flip = solves \word_eqI simp: simp simp_del: simp_del split: split split_del: split_del cong: cong simp flip: flip; (fastforce dest: test_bit_size simp: word_eqI_simps simp flip: flip simp: simp simp del: simp_del split: split split del: split_del cong: cong)?\ end diff --git a/thys/Word_Lib/Word_Lemmas.thy b/thys/Word_Lib/Word_Lemmas.thy --- a/thys/Word_Lib/Word_Lemmas.thy +++ b/thys/Word_Lib/Word_Lemmas.thy @@ -1,6213 +1,6212 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Lemmas with Generic Word Length" theory Word_Lemmas imports Word_EqI Word_Enum "HOL-Library.Sublist" begin lemma word_plus_mono_left: fixes x :: "'a :: len word" shows "\y \ z; x \ x + z\ \ y + x \ z + x" by unat_arith lemma word_shiftl_add_distrib: fixes x :: "'a :: len word" shows "(x + y) << n = (x << n) + (y << n)" by (simp add: shiftl_t2n ring_distribs) lemma less_Suc_unat_less_bound: "n < Suc (unat (x :: 'a :: len word)) \ n < 2 ^ LENGTH('a)" by (auto elim!: order_less_le_trans intro: Suc_leI) lemma up_ucast_inj: "\ ucast x = (ucast y::'b::len word); LENGTH('a) \ len_of TYPE ('b) \ \ x = (y::'a::len word)" by (subst (asm) bang_eq) (fastforce simp: nth_ucast word_size intro: word_eqI) lemmas ucast_up_inj = up_ucast_inj lemma up_ucast_inj_eq: "LENGTH('a) \ len_of TYPE ('b) \ (ucast x = (ucast y::'b::len word)) = (x = (y::'a::len word))" by (fastforce dest: up_ucast_inj) lemma no_plus_overflow_neg: "(x :: 'a :: len word) < -y \ x \ x + y" by (metis diff_minus_eq_add less_imp_le sub_wrap_lt) lemma ucast_ucast_eq: "\ ucast x = (ucast (ucast y::'a word)::'c::len word); LENGTH('a) \ LENGTH('b); LENGTH('b) \ LENGTH('c) \ \ x = ucast y" for x :: "'a::len word" and y :: "'b::len word" by (fastforce intro: word_eqI simp: bang_eq nth_ucast word_size) lemma ucast_0_I: "x = 0 \ ucast x = 0" by simp text \right-padding a word to a certain length\ definition "bl_pad_to bl sz \ bl @ (replicate (sz - length bl) False)" lemma bl_pad_to_length: assumes lbl: "length bl \ sz" shows "length (bl_pad_to bl sz) = sz" using lbl by (simp add: bl_pad_to_def) lemma bl_pad_to_prefix: "prefix bl (bl_pad_to bl sz)" by (simp add: bl_pad_to_def) lemma same_length_is_parallel: assumes len: "\y \ set as. length y = x" shows "\x \ set as. \y \ set as - {x}. x \ y" proof (rule, rule) fix x y assume xi: "x \ set as" and yi: "y \ set as - {x}" from len obtain q where len': "\y \ set as. length y = q" .. show "x \ y" proof (rule not_equal_is_parallel) from xi yi show "x \ y" by auto from xi yi len' show "length x = length y" by (auto dest: bspec) qed qed text \Lemmas about words\ lemmas and_bang = word_and_nth lemma of_drop_to_bl: "of_bl (drop n (to_bl x)) = (x && mask (size x - n))" by (simp add: of_bl_drop word_size_bl) lemma word_add_offset_less: fixes x :: "'a :: len word" assumes yv: "y < 2 ^ n" and xv: "x < 2 ^ m" and mnv: "sz < LENGTH('a :: len)" and xv': "x < 2 ^ (LENGTH('a :: len) - n)" and mn: "sz = m + n" shows "x * 2 ^ n + y < 2 ^ sz" proof (subst mn) from mnv mn have nv: "n < LENGTH('a)" and mv: "m < LENGTH('a)" by auto have uy: "unat y < 2 ^ n" by (rule order_less_le_trans [OF unat_mono [OF yv] order_eq_refl], rule unat_power_lower[OF nv]) have ux: "unat x < 2 ^ m" by (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl], rule unat_power_lower[OF mv]) then show "x * 2 ^ n + y < 2 ^ (m + n)" using ux uy nv mnv xv' apply (subst word_less_nat_alt) apply (subst unat_word_ariths)+ apply (subst mod_less) apply simp apply (subst mult.commute) apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]]) apply (rule order_less_le_trans [OF unat_mono [OF xv']]) apply (cases "n = 0"; simp) apply (subst unat_power_lower[OF nv]) apply (subst mod_less) apply (erule order_less_le_trans [OF nat_add_offset_less], assumption) apply (rule mn) apply simp apply (simp add: mn mnv) apply (erule nat_add_offset_less; simp) done qed lemma word_less_power_trans: fixes n :: "'a :: len word" assumes nv: "n < 2 ^ (m - k)" and kv: "k \ m" and mv: "m < len_of TYPE ('a)" shows "2 ^ k * n < 2 ^ m" using nv kv mv apply - apply (subst word_less_nat_alt) apply (subst unat_word_ariths) apply (subst mod_less) apply simp apply (rule nat_less_power_trans) apply (erule order_less_trans [OF unat_mono]) apply simp apply simp apply simp apply (rule nat_less_power_trans) apply (subst unat_power_lower[where 'a = 'a, symmetric]) apply simp apply (erule unat_mono) apply simp done lemma Suc_unat_diff_1: fixes x :: "'a :: len word" assumes lt: "1 \ x" shows "Suc (unat (x - 1)) = unat x" proof - have "0 < unat x" by (rule order_less_le_trans [where y = 1], simp, subst unat_1 [symmetric], rule iffD1 [OF word_le_nat_alt lt]) then show ?thesis by ((subst unat_sub [OF lt])+, simp only: unat_1) qed lemma word_div_sub: fixes x :: "'a :: len word" assumes yx: "y \ x" and y0: "0 < y" shows "(x - y) div y = x div y - 1" apply (rule word_unat.Rep_eqD) apply (subst unat_div) apply (subst unat_sub [OF yx]) apply (subst unat_sub) apply (subst word_le_nat_alt) apply (subst unat_div) apply (subst le_div_geq) apply (rule order_le_less_trans [OF _ unat_mono [OF y0]]) apply simp apply (subst word_le_nat_alt [symmetric], rule yx) apply simp apply (subst unat_div) apply (subst le_div_geq [OF _ iffD1 [OF word_le_nat_alt yx]]) apply (rule order_le_less_trans [OF _ unat_mono [OF y0]]) apply simp apply simp done lemma word_mult_less_mono1: fixes i :: "'a :: len word" assumes ij: "i < j" and knz: "0 < k" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i * k < j * k" proof - from ij ujk knz have jk: "unat i * unat k < 2 ^ len_of TYPE ('a)" by (auto intro: order_less_subst2 simp: word_less_nat_alt elim: mult_less_mono1) then show ?thesis using ujk knz ij by (auto simp: word_less_nat_alt iffD1 [OF unat_mult_lem]) qed lemma word_mult_less_dest: fixes i :: "'a :: len word" assumes ij: "i * k < j * k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i < j" using uik ujk ij by (auto simp: word_less_nat_alt iffD1 [OF unat_mult_lem] elim: mult_less_mono1) lemma word_mult_less_cancel: fixes k :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k < j * k) = (i < j)" by (rule iffI [OF word_mult_less_dest [OF _ uik ujk] word_mult_less_mono1 [OF _ knz ujk]]) lemma Suc_div_unat_helper: assumes szv: "sz < LENGTH('a :: len)" and usszv: "us \ sz" shows "2 ^ (sz - us) = Suc (unat (((2::'a :: len word) ^ sz - 1) div 2 ^ us))" proof - note usv = order_le_less_trans [OF usszv szv] from usszv obtain q where qv: "sz = us + q" by (auto simp: le_iff_add) have "Suc (unat (((2:: 'a word) ^ sz - 1) div 2 ^ us)) = (2 ^ us + unat ((2:: 'a word) ^ sz - 1)) div 2 ^ us" apply (subst unat_div unat_power_lower[OF usv])+ apply (subst div_add_self1, simp+) done also have "\ = ((2 ^ us - 1) + 2 ^ sz) div 2 ^ us" using szv by (simp add: unat_minus_one) also have "\ = 2 ^ q + ((2 ^ us - 1) div 2 ^ us)" apply (subst qv) apply (subst power_add) apply (subst div_mult_self2; simp) done also have "\ = 2 ^ (sz - us)" using qv by simp finally show ?thesis .. qed lemma set_enum_word8_def: "(set enum::word8 set) = {0, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21, 22, 23, 24, 25, 26, 27, 28, 29, 30, 31, 32, 33, 34, 35, 36, 37, 38, 39, 40, 41, 42, 43, 44, 45, 46, 47, 48, 49, 50, 51, 52, 53, 54, 55, 56, 57, 58, 59, 60, 61, 62, 63, 64, 65, 66, 67, 68, 69, 70, 71, 72, 73, 74, 75, 76, 77, 78, 79, 80, 81, 82, 83, 84, 85, 86, 87, 88, 89, 90, 91, 92, 93, 94, 95, 96, 97, 98, 99, 100, 101, 102, 103, 104, 105, 106, 107, 108, 109, 110, 111, 112, 113, 114, 115, 116, 117, 118, 119, 120, 121, 122, 123, 124, 125, 126, 127, 128, 129, 130, 131, 132, 133, 134, 135, 136, 137, 138, 139, 140, 141, 142, 143, 144, 145, 146, 147, 148, 149, 150, 151, 152, 153, 154, 155, 156, 157, 158, 159, 160, 161, 162, 163, 164, 165, 166, 167, 168, 169, 170, 171, 172, 173, 174, 175, 176, 177, 178, 179, 180, 181, 182, 183, 184, 185, 186, 187, 188, 189, 190, 191, 192, 193, 194, 195, 196, 197, 198, 199, 200, 201, 202, 203, 204, 205, 206, 207, 208, 209, 210, 211, 212, 213, 214, 215, 216, 217, 218, 219, 220, 221, 222, 223, 224, 225, 226, 227, 228, 229, 230, 231, 232, 233, 234, 235, 236, 237, 238, 239, 240, 241, 242, 243, 244, 245, 246, 247, 248, 249, 250, 251, 252, 253, 254, 255}" by eval lemma set_strip_insert: "\ x \ insert a S; x \ a \ \ x \ S" by simp lemma word8_exhaust: fixes x :: word8 shows "\x \ 0; x \ 1; x \ 2; x \ 3; x \ 4; x \ 5; x \ 6; x \ 7; x \ 8; x \ 9; x \ 10; x \ 11; x \ 12; x \ 13; x \ 14; x \ 15; x \ 16; x \ 17; x \ 18; x \ 19; x \ 20; x \ 21; x \ 22; x \ 23; x \ 24; x \ 25; x \ 26; x \ 27; x \ 28; x \ 29; x \ 30; x \ 31; x \ 32; x \ 33; x \ 34; x \ 35; x \ 36; x \ 37; x \ 38; x \ 39; x \ 40; x \ 41; x \ 42; x \ 43; x \ 44; x \ 45; x \ 46; x \ 47; x \ 48; x \ 49; x \ 50; x \ 51; x \ 52; x \ 53; x \ 54; x \ 55; x \ 56; x \ 57; x \ 58; x \ 59; x \ 60; x \ 61; x \ 62; x \ 63; x \ 64; x \ 65; x \ 66; x \ 67; x \ 68; x \ 69; x \ 70; x \ 71; x \ 72; x \ 73; x \ 74; x \ 75; x \ 76; x \ 77; x \ 78; x \ 79; x \ 80; x \ 81; x \ 82; x \ 83; x \ 84; x \ 85; x \ 86; x \ 87; x \ 88; x \ 89; x \ 90; x \ 91; x \ 92; x \ 93; x \ 94; x \ 95; x \ 96; x \ 97; x \ 98; x \ 99; x \ 100; x \ 101; x \ 102; x \ 103; x \ 104; x \ 105; x \ 106; x \ 107; x \ 108; x \ 109; x \ 110; x \ 111; x \ 112; x \ 113; x \ 114; x \ 115; x \ 116; x \ 117; x \ 118; x \ 119; x \ 120; x \ 121; x \ 122; x \ 123; x \ 124; x \ 125; x \ 126; x \ 127; x \ 128; x \ 129; x \ 130; x \ 131; x \ 132; x \ 133; x \ 134; x \ 135; x \ 136; x \ 137; x \ 138; x \ 139; x \ 140; x \ 141; x \ 142; x \ 143; x \ 144; x \ 145; x \ 146; x \ 147; x \ 148; x \ 149; x \ 150; x \ 151; x \ 152; x \ 153; x \ 154; x \ 155; x \ 156; x \ 157; x \ 158; x \ 159; x \ 160; x \ 161; x \ 162; x \ 163; x \ 164; x \ 165; x \ 166; x \ 167; x \ 168; x \ 169; x \ 170; x \ 171; x \ 172; x \ 173; x \ 174; x \ 175; x \ 176; x \ 177; x \ 178; x \ 179; x \ 180; x \ 181; x \ 182; x \ 183; x \ 184; x \ 185; x \ 186; x \ 187; x \ 188; x \ 189; x \ 190; x \ 191; x \ 192; x \ 193; x \ 194; x \ 195; x \ 196; x \ 197; x \ 198; x \ 199; x \ 200; x \ 201; x \ 202; x \ 203; x \ 204; x \ 205; x \ 206; x \ 207; x \ 208; x \ 209; x \ 210; x \ 211; x \ 212; x \ 213; x \ 214; x \ 215; x \ 216; x \ 217; x \ 218; x \ 219; x \ 220; x \ 221; x \ 222; x \ 223; x \ 224; x \ 225; x \ 226; x \ 227; x \ 228; x \ 229; x \ 230; x \ 231; x \ 232; x \ 233; x \ 234; x \ 235; x \ 236; x \ 237; x \ 238; x \ 239; x \ 240; x \ 241; x \ 242; x \ 243; x \ 244; x \ 245; x \ 246; x \ 247; x \ 248; x \ 249; x \ 250; x \ 251; x \ 252; x \ 253; x \ 254; x \ 255\ \ P" apply (subgoal_tac "x \ set enum", subst (asm) set_enum_word8_def) apply (drule set_strip_insert, assumption)+ apply (erule emptyE) apply (subst enum_UNIV, rule UNIV_I) done lemma upto_enum_red': assumes lt: "1 \ X" shows "[(0::'a :: len word) .e. X - 1] = map of_nat [0 ..< unat X]" proof - have lt': "unat X < 2 ^ LENGTH('a)" by (rule unat_lt2p) show ?thesis apply (subst upto_enum_red) apply (simp del: upt.simps) apply (subst Suc_unat_diff_1 [OF lt]) apply (rule map_cong [OF refl]) apply (rule toEnum_of_nat) apply simp apply (erule order_less_trans [OF _ lt']) done qed lemma upto_enum_red2: assumes szv: "sz < LENGTH('a :: len)" shows "[(0:: 'a :: len word) .e. 2 ^ sz - 1] = map of_nat [0 ..< 2 ^ sz]" using szv apply (subst unat_power_lower[OF szv, symmetric]) apply (rule upto_enum_red') apply (subst word_le_nat_alt, simp) done lemma upto_enum_step_red: assumes szv: "sz < LENGTH('a)" and usszv: "us \ sz" shows "[0 :: 'a :: len word , 2 ^ us .e. 2 ^ sz - 1] = map (\x. of_nat x * 2 ^ us) [0 ..< 2 ^ (sz - us)]" using szv unfolding upto_enum_step_def apply (subst if_not_P) apply (rule leD) apply (subst word_le_nat_alt) apply (subst unat_minus_one) apply simp apply simp apply simp apply (subst upto_enum_red) apply (simp del: upt.simps) apply (subst Suc_div_unat_helper [where 'a = 'a, OF szv usszv, symmetric]) apply clarsimp apply (subst toEnum_of_nat) apply (erule order_less_trans) using szv apply simp apply simp done lemma upto_enum_word: "[x .e. y] = map of_nat [unat x ..< Suc (unat y)]" apply (subst upto_enum_red) apply clarsimp apply (subst toEnum_of_nat) prefer 2 apply (rule refl) apply (erule disjE, simp) apply clarsimp apply (erule order_less_trans) apply simp done lemma word_upto_Cons_eq: "x < y \ [x::'a::len word .e. y] = x # [x + 1 .e. y]" apply (subst upto_enum_red) apply (subst upt_conv_Cons, unat_arith) apply (simp only: list.map list.inject upto_enum_red to_from_enum simp_thms) apply (rule map_cong[OF _ refl]) apply (rule arg_cong2[where f = "\x y. [x ..< y]"], unat_arith) apply (rule refl) done lemma distinct_enum_upto: "distinct [(0 :: 'a::len word) .e. b]" proof - have "\(b::'a word). [0 .e. b] = nths enum {..< Suc (fromEnum b)}" apply (subst upto_enum_red) apply (subst nths_upt_eq_take) apply (subst enum_word_def) apply (subst take_map) apply (subst take_upt) apply (simp only: add_0 fromEnum_unat) apply (rule order_trans [OF _ order_eq_refl]) apply (rule Suc_leI [OF unat_lt2p]) apply simp apply clarsimp apply (rule toEnum_of_nat) apply (erule order_less_trans [OF _ unat_lt2p]) done then show ?thesis by (rule ssubst) (rule distinct_nthsI, simp) qed lemma upto_enum_set_conv [simp]: fixes a :: "'a :: len word" shows "set [a .e. b] = {x. a \ x \ x \ b}" apply (subst upto_enum_red) apply (subst set_map) apply safe apply simp apply clarsimp apply (erule disjE) apply simp apply (erule iffD2 [OF word_le_nat_alt]) apply clarsimp apply (erule word_unat.Rep_cases [OF unat_le [OF order_less_imp_le]]) apply simp apply (erule iffD2 [OF word_le_nat_alt]) apply simp apply clarsimp apply (erule disjE) apply simp apply clarsimp apply (rule word_unat.Rep_cases [OF unat_le [OF order_less_imp_le]]) apply assumption apply simp apply (erule order_less_imp_le [OF iffD2 [OF word_less_nat_alt]]) apply clarsimp apply (rule_tac x="fromEnum x" in image_eqI) apply clarsimp apply clarsimp apply (rule conjI) apply (subst word_le_nat_alt [symmetric]) apply simp apply safe apply (simp add: word_le_nat_alt [symmetric]) apply (simp add: word_less_nat_alt [symmetric]) done lemma upto_enum_less: assumes xin: "x \ set [(a::'a::len word).e.2 ^ n - 1]" and nv: "n < LENGTH('a::len)" shows "x < 2 ^ n" proof (cases n) case 0 then show ?thesis using xin by simp next case (Suc m) show ?thesis using xin nv by simp qed lemma upto_enum_len_less: "\ n \ length [a, b .e. c]; n \ 0 \ \ a \ c" unfolding upto_enum_step_def by (simp split: if_split_asm) lemma length_upto_enum_step: fixes x :: "'a :: len word" shows "x \ z \ length [x , y .e. z] = (unat ((z - x) div (y - x))) + 1" unfolding upto_enum_step_def by (simp add: upto_enum_red) lemma map_length_unfold_one: fixes x :: "'a::len word" assumes xv: "Suc (unat x) < 2 ^ LENGTH('a)" and ax: "a < x" shows "map f [a .e. x] = f a # map f [a + 1 .e. x]" by (subst word_upto_Cons_eq, auto, fact+) lemma upto_enum_set_conv2: fixes a :: "'a::len word" shows "set [a .e. b] = {a .. b}" by auto lemma of_nat_unat [simp]: "of_nat \ unat = id" by (rule ext, simp) lemma Suc_unat_minus_one [simp]: "x \ 0 \ Suc (unat (x - 1)) = unat x" by (metis Suc_diff_1 unat_gt_0 unat_minus_one) lemma word_add_le_dest: fixes i :: "'a :: len word" assumes le: "i + k \ j + k" and uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i \ j" using uik ujk le by (auto simp: word_le_nat_alt iffD1 [OF unat_add_lem] elim: add_le_mono1) lemma mask_shift: "(x && ~~ (mask y)) >> y = x >> y" by word_eqI lemma word_add_le_mono1: fixes i :: "'a :: len word" assumes ij: "i \ j" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i + k \ j + k" proof - from ij ujk have jk: "unat i + unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_le_nat_alt elim: add_le_mono1) then show ?thesis using ujk ij by (auto simp: word_le_nat_alt iffD1 [OF unat_add_lem]) qed lemma word_add_le_mono2: fixes i :: "'a :: len word" shows "\i \ j; unat j + unat k < 2 ^ LENGTH('a)\ \ k + i \ k + j" by (subst field_simps, subst field_simps, erule (1) word_add_le_mono1) lemma word_add_le_iff: fixes i :: "'a :: len word" assumes uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "(i + k \ j + k) = (i \ j)" proof assume "i \ j" show "i + k \ j + k" by (rule word_add_le_mono1) fact+ next assume "i + k \ j + k" show "i \ j" by (rule word_add_le_dest) fact+ qed lemma word_add_less_mono1: fixes i :: "'a :: len word" assumes ij: "i < j" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i + k < j + k" proof - from ij ujk have jk: "unat i + unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_less_nat_alt elim: add_less_mono1) then show ?thesis using ujk ij by (auto simp: word_less_nat_alt iffD1 [OF unat_add_lem]) qed lemma word_add_less_dest: fixes i :: "'a :: len word" assumes le: "i + k < j + k" and uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i < j" using uik ujk le by (auto simp: word_less_nat_alt iffD1 [OF unat_add_lem] elim: add_less_mono1) lemma word_add_less_iff: fixes i :: "'a :: len word" assumes uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "(i + k < j + k) = (i < j)" proof assume "i < j" show "i + k < j + k" by (rule word_add_less_mono1) fact+ next assume "i + k < j + k" show "i < j" by (rule word_add_less_dest) fact+ qed lemma shiftr_div_2n': "unat (w >> n) = unat w div 2 ^ n" apply (unfold unat_def) apply (subst shiftr_div_2n) apply (subst nat_div_distrib) apply simp apply (simp add: nat_power_eq) done lemma shiftl_shiftr_id: assumes nv: "n < LENGTH('a)" and xv: "x < 2 ^ (LENGTH('a) - n)" shows "x << n >> n = (x::'a::len word)" apply (simp add: shiftl_t2n) apply (rule word_unat.Rep_eqD) apply (subst shiftr_div_2n') apply (cases n) apply simp apply (subst iffD1 [OF unat_mult_lem])+ apply (subst unat_power_lower[OF nv]) apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]]) apply (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl]) apply (rule unat_power_lower) apply simp apply (subst unat_power_lower[OF nv]) apply simp done lemma ucast_shiftl_eq_0: fixes w :: "'a :: len word" shows "\ n \ LENGTH('b) \ \ ucast (w << n) = (0 :: 'b :: len word)" by (case_tac "size w \ n", clarsimp simp: shiftl_zero_size) (clarsimp simp: not_le ucast_bl bl_shiftl bang_eq test_bit_of_bl rev_nth nth_append) lemma word_mult_less_iff: fixes i :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k < j * k) = (i < j)" using assms by (rule word_mult_less_cancel) lemma word_le_imp_diff_le: fixes n :: "'a::len word" shows "\k \ n; n \ m\ \ n - k \ m" by (auto simp: unat_sub word_le_nat_alt) lemma word_less_imp_diff_less: fixes n :: "'a::len word" shows "\k \ n; n < m\ \ n - k < m" by (clarsimp simp: unat_sub word_less_nat_alt intro!: less_imp_diff_less) lemma word_mult_le_mono1: fixes i :: "'a :: len word" assumes ij: "i \ j" and knz: "0 < k" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i * k \ j * k" proof - from ij ujk knz have jk: "unat i * unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_le_nat_alt elim: mult_le_mono1) then show ?thesis using ujk knz ij by (auto simp: word_le_nat_alt iffD1 [OF unat_mult_lem]) qed lemma word_mult_le_iff: fixes i :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k \ j * k) = (i \ j)" proof assume "i \ j" show "i * k \ j * k" by (rule word_mult_le_mono1) fact+ next assume p: "i * k \ j * k" have "0 < unat k" using knz by (simp add: word_less_nat_alt) then show "i \ j" using p by (clarsimp simp: word_le_nat_alt iffD1 [OF unat_mult_lem uik] iffD1 [OF unat_mult_lem ujk]) qed lemma word_diff_less: fixes n :: "'a :: len word" shows "\0 < n; 0 < m; n \ m\ \ m - n < m" apply (subst word_less_nat_alt) apply (subst unat_sub) apply assumption apply (rule diff_less) apply (simp_all add: word_less_nat_alt) done lemma MinI: assumes fa: "finite A" and ne: "A \ {}" and xv: "m \ A" and min: "\y \ A. m \ y" shows "Min A = m" using fa ne xv min proof (induct A arbitrary: m rule: finite_ne_induct) case singleton then show ?case by simp next case (insert y F) from insert.prems have yx: "m \ y" and fx: "\y \ F. m \ y" by auto have "m \ insert y F" by fact then show ?case proof assume mv: "m = y" have mlt: "m \ Min F" by (rule iffD2 [OF Min_ge_iff [OF insert.hyps(1) insert.hyps(2)] fx]) show ?case apply (subst Min_insert [OF insert.hyps(1) insert.hyps(2)]) apply (subst mv [symmetric]) apply (auto simp: min_def mlt) done next assume "m \ F" then have mf: "Min F = m" by (rule insert.hyps(4) [OF _ fx]) show ?case apply (subst Min_insert [OF insert.hyps(1) insert.hyps(2)]) apply (subst mf) apply (rule iffD2 [OF _ yx]) apply (auto simp: min_def) done qed qed lemma length_upto_enum [simp]: fixes a :: "'a :: len word" shows "length [a .e. b] = Suc (unat b) - unat a" apply (simp add: word_le_nat_alt upto_enum_red) apply (clarsimp simp: Suc_diff_le) done lemma length_upto_enum_cases: fixes a :: "'a::len word" shows "length [a .e. b] = (if a \ b then Suc (unat b) - unat a else 0)" apply (case_tac "a \ b") apply (clarsimp) apply (clarsimp simp: upto_enum_def) apply unat_arith done lemma length_upto_enum_less_one: "\a \ b; b \ 0\ \ length [a .e. b - 1] = unat (b - a)" apply clarsimp apply (subst unat_sub[symmetric], assumption) apply clarsimp done lemma drop_upto_enum: "drop (unat n) [0 .e. m] = [n .e. m]" apply (clarsimp simp: upto_enum_def) apply (induct m, simp) by (metis drop_map drop_upt plus_nat.add_0) lemma distinct_enum_upto' [simp]: "distinct [a::'a::len word .e. b]" apply (subst drop_upto_enum [symmetric]) apply (rule distinct_drop) apply (rule distinct_enum_upto) done lemma length_interval: "\set xs = {x. (a::'a::len word) \ x \ x \ b}; distinct xs\ \ length xs = Suc (unat b) - unat a" apply (frule distinct_card) apply (subgoal_tac "set xs = set [a .e. b]") apply (cut_tac distinct_card [where xs="[a .e. b]"]) apply (subst (asm) length_upto_enum) apply clarsimp apply (rule distinct_enum_upto') apply simp done lemma not_empty_eq: "(S \ {}) = (\x. x \ S)" by auto lemma range_subset_lower: fixes c :: "'a ::linorder" shows "\ {a..b} \ {c..d}; x \ {a..b} \ \ c \ a" apply (frule (1) subsetD) apply (rule classical) apply clarsimp done lemma range_subset_upper: fixes c :: "'a ::linorder" shows "\ {a..b} \ {c..d}; x \ {a..b} \ \ b \ d" apply (frule (1) subsetD) apply (rule classical) apply clarsimp done lemma range_subset_eq: fixes a::"'a::linorder" assumes non_empty: "a \ b" shows "({a..b} \ {c..d}) = (c \ a \ b \ d)" apply (insert non_empty) apply (rule iffI) apply (frule range_subset_lower [where x=a], simp) apply (drule range_subset_upper [where x=a], simp) apply simp apply auto done lemma range_eq: fixes a::"'a::linorder" assumes non_empty: "a \ b" shows "({a..b} = {c..d}) = (a = c \ b = d)" by (metis atLeastatMost_subset_iff eq_iff non_empty) lemma range_strict_subset_eq: fixes a::"'a::linorder" assumes non_empty: "a \ b" shows "({a..b} \ {c..d}) = (c \ a \ b \ d \ (a = c \ b \ d))" apply (insert non_empty) apply (subst psubset_eq) apply (subst range_subset_eq, assumption+) apply (subst range_eq, assumption+) apply simp done lemma range_subsetI: fixes x :: "'a :: order" assumes xX: "X \ x" and yY: "y \ Y" shows "{x .. y} \ {X .. Y}" using xX yY by auto lemma set_False [simp]: "(set bs \ {False}) = (True \ set bs)" by auto declare of_nat_power [simp del] (* TODO: move to word *) lemma unat_of_bl_length: "unat (of_bl xs :: 'a::len word) < 2 ^ (length xs)" proof (cases "length xs < LENGTH('a)") case True then have "(of_bl xs::'a::len word) < 2 ^ length xs" by (simp add: of_bl_length_less) with True show ?thesis by (simp add: word_less_nat_alt word_unat_power unat_of_nat) next case False have "unat (of_bl xs::'a::len word) < 2 ^ LENGTH('a)" by (simp split: unat_split) also from False have "LENGTH('a) \ length xs" by simp then have "2 ^ LENGTH('a) \ (2::nat) ^ length xs" by (rule power_increasing) simp finally show ?thesis . qed lemma is_aligned_0'[simp]: "is_aligned 0 n" by (simp add: is_aligned_def) lemma p_assoc_help: fixes p :: "'a::{ring,power,numeral,one}" shows "p + 2^sz - 1 = p + (2^sz - 1)" by simp lemma word_add_increasing: fixes x :: "'a :: len word" shows "\ p + w \ x; p \ p + w \ \ p \ x" by unat_arith lemma word_random: fixes x :: "'a :: len word" shows "\ p \ p + x'; x \ x' \ \ p \ p + x" by unat_arith lemma word_sub_mono: "\ a \ c; d \ b; a - b \ a; c - d \ c \ \ (a - b) \ (c - d :: 'a :: len word)" by unat_arith lemma power_not_zero: "n < LENGTH('a::len) \ (2 :: 'a word) ^ n \ 0" by (metis p2_gt_0 word_neq_0_conv) lemma word_gt_a_gt_0: "a < n \ (0 :: 'a::len word) < n" apply (case_tac "n = 0") apply clarsimp apply (clarsimp simp: word_neq_0_conv) done lemma word_shift_nonzero: "\ (x::'a::len word) \ 2 ^ m; m + n < LENGTH('a::len); x \ 0\ \ x << n \ 0" apply (simp only: word_neq_0_conv word_less_nat_alt shiftl_t2n mod_0 unat_word_ariths unat_power_lower word_le_nat_alt) apply (subst mod_less) apply (rule order_le_less_trans) apply (erule mult_le_mono2) apply (subst power_add[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply simp done lemma word_power_less_1 [simp]: "sz < LENGTH('a::len) \ (2::'a word) ^ sz - 1 < 2 ^ sz" apply (simp add: word_less_nat_alt) apply (subst unat_minus_one) apply (simp add: word_unat.Rep_inject [symmetric]) apply simp done lemma nasty_split_lt: "\ (x :: 'a:: len word) < 2 ^ (m - n); n \ m; m < LENGTH('a::len) \ \ x * 2 ^ n + (2 ^ n - 1) \ 2 ^ m - 1" apply (simp only: add_diff_eq) apply (subst mult_1[symmetric], subst distrib_right[symmetric]) apply (rule word_sub_mono) apply (rule order_trans) apply (rule word_mult_le_mono1) apply (rule inc_le) apply assumption apply (subst word_neq_0_conv[symmetric]) apply (rule power_not_zero) apply simp apply (subst unat_power_lower, simp)+ apply (subst power_add[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply (subst power_add[symmetric]) apply simp apply simp apply (rule word_sub_1_le) apply (subst mult.commute) apply (subst shiftl_t2n[symmetric]) apply (rule word_shift_nonzero) apply (erule inc_le) apply simp apply (unat_arith) apply (drule word_power_less_1) apply simp done lemma nasty_split_less: "\m \ n; n \ nm; nm < LENGTH('a::len); x < 2 ^ (nm - n)\ \ (x :: 'a word) * 2 ^ n + (2 ^ m - 1) < 2 ^ nm" apply (simp only: word_less_sub_le[symmetric]) apply (rule order_trans [OF _ nasty_split_lt]) apply (rule word_plus_mono_right) apply (rule word_sub_mono) apply (simp add: word_le_nat_alt) apply simp apply (simp add: word_sub_1_le[OF power_not_zero]) apply (simp add: word_sub_1_le[OF power_not_zero]) apply (rule is_aligned_no_wrap') apply (rule is_aligned_mult_triv2) apply simp apply (erule order_le_less_trans, simp) apply simp+ done lemma int_not_emptyD: "A \ B \ {} \ \x. x \ A \ x \ B" by (erule contrapos_np, clarsimp simp: disjoint_iff_not_equal) lemma unat_less_power: fixes k :: "'a::len word" assumes szv: "sz < LENGTH('a)" and kv: "k < 2 ^ sz" shows "unat k < 2 ^ sz" using szv unat_mono [OF kv] by simp lemma unat_mult_power_lem: assumes kv: "k < 2 ^ (LENGTH('a::len) - sz)" shows "unat (2 ^ sz * of_nat k :: (('a::len) word)) = 2 ^ sz * k" proof cases assume szv: "sz < LENGTH('a::len)" show ?thesis proof (cases "sz = 0") case True then show ?thesis using kv szv by (simp add: unat_of_nat) next case False then have sne: "0 < sz" .. have uk: "unat (of_nat k :: 'a word) = k" apply (subst unat_of_nat) apply (simp add: nat_mod_eq less_trans[OF kv] sne) done show ?thesis using szv apply (subst iffD1 [OF unat_mult_lem]) apply (simp add: uk nat_less_power_trans[OF kv order_less_imp_le [OF szv]])+ done qed next assume "\ sz < LENGTH('a)" with kv show ?thesis by (simp add: not_less power_overflow) qed lemma aligned_add_offset_no_wrap: fixes off :: "('a::len) word" and x :: "'a word" assumes al: "is_aligned x sz" and offv: "off < 2 ^ sz" shows "unat x + unat off < 2 ^ LENGTH('a)" proof cases assume szv: "sz < LENGTH('a)" from al obtain k where xv: "x = 2 ^ sz * (of_nat k)" and kl: "k < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) show ?thesis using szv apply (subst xv) apply (subst unat_mult_power_lem[OF kl]) apply (subst mult.commute, rule nat_add_offset_less) apply (rule less_le_trans[OF unat_mono[OF offv, simplified]]) apply (erule eq_imp_le[OF unat_power_lower]) apply (rule kl) apply simp done next assume "\ sz < LENGTH('a)" with offv show ?thesis by (simp add: not_less power_overflow ) qed lemma aligned_add_offset_mod: fixes x :: "('a::len) word" assumes al: "is_aligned x sz" and kv: "k < 2 ^ sz" shows "(x + k) mod 2 ^ sz = k" proof cases assume szv: "sz < LENGTH('a)" have ux: "unat x + unat k < 2 ^ LENGTH('a)" by (rule aligned_add_offset_no_wrap) fact+ show ?thesis using al szv apply - apply (erule is_alignedE) apply (subst word_unat.Rep_inject [symmetric]) apply (subst unat_mod) apply (subst iffD1 [OF unat_add_lem], rule ux) apply simp apply (subst unat_mult_power_lem, assumption+) apply (simp) apply (rule mod_less[OF less_le_trans[OF unat_mono], OF kv]) apply (erule eq_imp_le[OF unat_power_lower]) done next assume "\ sz < LENGTH('a)" with al show ?thesis by (simp add: not_less power_overflow is_aligned_mask mask_def word_mod_by_0) qed lemma word_plus_mcs_4: "\v + x \ w + x; x \ v + x\ \ v \ (w::'a::len word)" by uint_arith lemma word_plus_mcs_3: "\v \ w; x \ w + x\ \ v + x \ w + (x::'a::len word)" by unat_arith lemma aligned_neq_into_no_overlap: fixes x :: "'a::len word" assumes neq: "x \ y" and alx: "is_aligned x sz" and aly: "is_aligned y sz" shows "{x .. x + (2 ^ sz - 1)} \ {y .. y + (2 ^ sz - 1)} = {}" proof cases assume szv: "sz < LENGTH('a)" show ?thesis proof (rule equals0I, clarsimp) fix z assume xb: "x \ z" and xt: "z \ x + (2 ^ sz - 1)" and yb: "y \ z" and yt: "z \ y + (2 ^ sz - 1)" have rl: "\(p::'a word) k w. \uint p + uint k < 2 ^ LENGTH('a); w = p + k; w \ p + (2 ^ sz - 1) \ \ k < 2 ^ sz" apply - apply simp apply (subst (asm) add.commute, subst (asm) add.commute, drule word_plus_mcs_4) apply (subst add.commute, subst no_plus_overflow_uint_size) apply (simp add: word_size_bl) apply (erule iffD1 [OF word_less_sub_le[OF szv]]) done from xb obtain kx where kx: "z = x + kx" and kxl: "uint x + uint kx < 2 ^ LENGTH('a)" by (clarsimp dest!: word_le_exists') from yb obtain ky where ky: "z = y + ky" and kyl: "uint y + uint ky < 2 ^ LENGTH('a)" by (clarsimp dest!: word_le_exists') have "x = y" proof - have "kx = z mod 2 ^ sz" proof (subst kx, rule sym, rule aligned_add_offset_mod) show "kx < 2 ^ sz" by (rule rl) fact+ qed fact+ also have "\ = ky" proof (subst ky, rule aligned_add_offset_mod) show "ky < 2 ^ sz" using kyl ky yt by (rule rl) qed fact+ finally have kxky: "kx = ky" . moreover have "x + kx = y + ky" by (simp add: kx [symmetric] ky [symmetric]) ultimately show ?thesis by simp qed then show False using neq by simp qed next assume "\ sz < LENGTH('a)" with neq alx aly have False by (simp add: is_aligned_mask mask_def power_overflow) then show ?thesis .. qed lemma less_two_pow_divD: "\ (x :: nat) < 2 ^ n div 2 ^ m \ \ n \ m \ (x < 2 ^ (n - m))" apply (rule context_conjI) apply (rule ccontr) apply (simp add: power_strict_increasing) apply (simp add: power_sub) done lemma less_two_pow_divI: "\ (x :: nat) < 2 ^ (n - m); m \ n \ \ x < 2 ^ n div 2 ^ m" by (simp add: power_sub) lemma word_less_two_pow_divI: "\ (x :: 'a::len word) < 2 ^ (n - m); m \ n; n < LENGTH('a) \ \ x < 2 ^ n div 2 ^ m" apply (simp add: word_less_nat_alt) apply (subst unat_word_ariths) apply (subst mod_less) apply (rule order_le_less_trans [OF div_le_dividend]) apply (rule unat_lt2p) apply (simp add: power_sub) done lemma word_less_two_pow_divD: "\ (x :: 'a::len word) < 2 ^ n div 2 ^ m \ \ n \ m \ (x < 2 ^ (n - m))" apply (cases "n < LENGTH('a)") apply (cases "m < LENGTH('a)") apply (simp add: word_less_nat_alt) apply (subst(asm) unat_word_ariths) apply (subst(asm) mod_less) apply (rule order_le_less_trans [OF div_le_dividend]) apply (rule unat_lt2p) apply (clarsimp dest!: less_two_pow_divD) apply (simp add: power_overflow) apply (simp add: word_div_def) apply (simp add: power_overflow word_div_def) done lemma of_nat_less_two_pow_div_set: "\ n < LENGTH('a) \ \ {x. x < (2 ^ n div 2 ^ m :: 'a::len word)} = of_nat ` {k. k < 2 ^ n div 2 ^ m}" apply (simp add: image_def) apply (safe dest!: word_less_two_pow_divD less_two_pow_divD intro!: word_less_two_pow_divI) apply (rule_tac x="unat x" in exI) apply (simp add: power_sub[symmetric]) apply (subst unat_power_lower[symmetric, where 'a='a]) apply simp apply (erule unat_mono) apply (subst word_unat_power) apply (rule of_nat_mono_maybe) apply (rule power_strict_increasing) apply simp apply simp apply assumption done lemma word_less_power_trans2: fixes n :: "'a::len word" shows "\n < 2 ^ (m - k); k \ m; m < LENGTH('a)\ \ n * 2 ^ k < 2 ^ m" by (subst field_simps, rule word_less_power_trans) (* shadows the slightly weaker Word.nth_ucast *) lemma nth_ucast: - "(ucast (w::'a::len0 word)::'b::len0 word) !! n = + "(ucast (w::'a::len word)::'b::len word) !! n = (w !! n \ n < min LENGTH('a) LENGTH('b))" by (simp add: ucast_def test_bit_bin word_ubin.eq_norm nth_bintr word_size) (fast elim!: bin_nth_uint_imp) lemma ucast_less: "LENGTH('b) < LENGTH('a) \ (ucast (x :: 'b :: len word) :: ('a :: len word)) < 2 ^ LENGTH('b)" by (meson Word.nth_ucast test_bit_conj_lt le_def upper_bits_unset_is_l2p) lemma ucast_range_less: "LENGTH('a :: len) < LENGTH('b :: len) \ range (ucast :: 'a word \ 'b word) = {x. x < 2 ^ len_of TYPE ('a)}" apply safe apply (erule ucast_less) apply (simp add: image_def) apply (rule_tac x="ucast x" in exI) by word_eqI_solve lemma word_power_less_diff: "\2 ^ n * q < (2::'a::len word) ^ m; q < 2 ^ (LENGTH('a) - n)\ \ q < 2 ^ (m - n)" apply (case_tac "m \ LENGTH('a)") apply (simp add: power_overflow) apply (case_tac "n \ LENGTH('a)") apply (simp add: power_overflow) apply (cases "n = 0") apply simp apply (subst word_less_nat_alt) apply (subst unat_power_lower) apply simp apply (rule nat_power_less_diff) apply (simp add: word_less_nat_alt) apply (subst (asm) iffD1 [OF unat_mult_lem]) apply (simp add:nat_less_power_trans) apply simp done lemmas word_diff_ls'' = word_diff_ls [where xa=x and x=x for x] lemmas word_diff_ls' = word_diff_ls'' [simplified] lemmas word_l_diffs' = word_l_diffs [where xa=x and x=x for x] lemmas word_l_diffs = word_l_diffs' [simplified] lemma is_aligned_diff: fixes m :: "'a::len word" assumes alm: "is_aligned m s1" and aln: "is_aligned n s2" and s2wb: "s2 < LENGTH('a)" and nm: "m \ {n .. n + (2 ^ s2 - 1)}" and s1s2: "s1 \ s2" and s10: "0 < s1" (* Probably can be folded into the proof \ *) shows "\q. m - n = of_nat q * 2 ^ s1 \ q < 2 ^ (s2 - s1)" proof - have rl: "\m s. \ m < 2 ^ (LENGTH('a) - s); s < LENGTH('a) \ \ unat ((2::'a word) ^ s * of_nat m) = 2 ^ s * m" proof - fix m :: nat and s assume m: "m < 2 ^ (LENGTH('a) - s)" and s: "s < LENGTH('a)" then have "unat ((of_nat m) :: 'a word) = m" apply (subst unat_of_nat) apply (subst mod_less) apply (erule order_less_le_trans) apply (rule power_increasing) apply simp_all done then show "?thesis m s" using s m apply (subst iffD1 [OF unat_mult_lem]) apply (simp add: nat_less_power_trans)+ done qed have s1wb: "s1 < LENGTH('a)" using s2wb s1s2 by simp from alm obtain mq where mmq: "m = 2 ^ s1 * of_nat mq" and mq: "mq < 2 ^ (LENGTH('a) - s1)" by (auto elim: is_alignedE simp: field_simps) from aln obtain nq where nnq: "n = 2 ^ s2 * of_nat nq" and nq: "nq < 2 ^ (LENGTH('a) - s2)" by (auto elim: is_alignedE simp: field_simps) from s1s2 obtain sq where sq: "s2 = s1 + sq" by (auto simp: le_iff_add) note us1 = rl [OF mq s1wb] note us2 = rl [OF nq s2wb] from nm have "n \ m" by clarsimp then have "(2::'a word) ^ s2 * of_nat nq \ 2 ^ s1 * of_nat mq" using nnq mmq by simp then have "2 ^ s2 * nq \ 2 ^ s1 * mq" using s1wb s2wb by (simp add: word_le_nat_alt us1 us2) then have nqmq: "2 ^ sq * nq \ mq" using sq by (simp add: power_add) have "m - n = 2 ^ s1 * of_nat mq - 2 ^ s2 * of_nat nq" using mmq nnq by simp also have "\ = 2 ^ s1 * of_nat mq - 2 ^ s1 * 2 ^ sq * of_nat nq" using sq by (simp add: power_add) also have "\ = 2 ^ s1 * (of_nat mq - 2 ^ sq * of_nat nq)" by (simp add: field_simps) also have "\ = 2 ^ s1 * of_nat (mq - 2 ^ sq * nq)" using s1wb s2wb us1 us2 nqmq by (simp add: word_unat_power) finally have mn: "m - n = of_nat (mq - 2 ^ sq * nq) * 2 ^ s1" by simp moreover from nm have "m - n \ 2 ^ s2 - 1" by - (rule word_diff_ls', (simp add: field_simps)+) then have "(2::'a word) ^ s1 * of_nat (mq - 2 ^ sq * nq) < 2 ^ s2" using mn s2wb by (simp add: field_simps) then have "of_nat (mq - 2 ^ sq * nq) < (2::'a word) ^ (s2 - s1)" proof (rule word_power_less_diff) have mm: "mq - 2 ^ sq * nq < 2 ^ (LENGTH('a) - s1)" using mq by simp moreover from s10 have "LENGTH('a) - s1 < LENGTH('a)" by (rule diff_less, simp) ultimately show "of_nat (mq - 2 ^ sq * nq) < (2::'a word) ^ (LENGTH('a) - s1)" apply (simp add: word_less_nat_alt) apply (subst unat_of_nat) apply (subst mod_less) apply (erule order_less_le_trans) apply simp+ done qed then have "mq - 2 ^ sq * nq < 2 ^ (s2 - s1)" using mq s2wb apply (simp add: word_less_nat_alt) apply (subst (asm) unat_of_nat) apply (subst (asm) mod_less) apply (rule order_le_less_trans) apply (rule diff_le_self) apply (erule order_less_le_trans) apply simp apply assumption done ultimately show ?thesis by auto qed lemma word_less_sub_1: "x < (y :: 'a :: len word) \ x \ y - 1" apply (erule udvd_minus_le') apply (simp add: udvd_def)+ done lemma word_sub_mono2: "\ a + b \ c + d; c \ a; b \ a + b; d \ c + d \ \ b \ (d :: 'a :: len word)" apply (drule(1) word_sub_mono) apply simp apply simp apply simp done lemma word_not_le: "(\ x \ (y :: 'a :: len word)) = (y < x)" by fastforce lemma word_subset_less: "\ {x .. x + r - 1} \ {y .. y + s - 1}; x \ x + r - 1; y \ y + (s :: 'a :: len word) - 1; s \ 0 \ \ r \ s" apply (frule subsetD[where c=x]) apply simp apply (drule subsetD[where c="x + r - 1"]) apply simp apply (clarsimp simp: add_diff_eq[symmetric]) apply (drule(1) word_sub_mono2) apply (simp_all add: olen_add_eqv[symmetric]) apply (erule word_le_minus_cancel) apply (rule ccontr) apply (simp add: word_not_le) done lemma uint_power_lower: "n < LENGTH('a) \ uint (2 ^ n :: 'a :: len word) = (2 ^ n :: int)" by (rule uint_2p_alt) lemma power_le_mono: "\2 ^ n \ (2::'a::len word) ^ m; n < LENGTH('a); m < LENGTH('a)\ \ n \ m" apply (clarsimp simp add: le_less) apply safe apply (simp add: word_less_nat_alt) apply (simp only: uint_arith_simps(3)) apply (drule uint_power_lower)+ apply simp done lemma sublist_equal_part: "prefix xs ys \ take (length xs) ys = xs" by (clarsimp simp: prefix_def) lemma two_power_eq: "\n < LENGTH('a); m < LENGTH('a)\ \ ((2::'a::len word) ^ n = 2 ^ m) = (n = m)" apply safe apply (rule order_antisym) apply (simp add: power_le_mono[where 'a='a])+ done lemma prefix_length_less: "strict_prefix xs ys \ length xs < length ys" apply (clarsimp simp: strict_prefix_def) apply (frule prefix_length_le) apply (rule ccontr, simp) apply (clarsimp simp: prefix_def) done lemmas take_less = take_strict_prefix lemma not_prefix_longer: "\ length xs > length ys \ \ \ prefix xs ys" by (clarsimp dest!: prefix_length_le) lemma of_bl_length: "length xs < LENGTH('a) \ of_bl xs < (2 :: 'a::len word) ^ length xs" by (simp add: of_bl_length_less) lemma unat_of_nat_eq: "x < 2 ^ LENGTH('a) \ unat (of_nat x ::'a::len word) = x" by (rule unat_of_nat_len) lemma unat_eq_of_nat: "n < 2 ^ LENGTH('a) \ (unat (x :: 'a::len word) = n) = (x = of_nat n)" by (subst unat_of_nat_eq[where x=n, symmetric], simp+) lemma unat_less_helper: "x < of_nat n \ unat x < n" apply (simp add: word_less_nat_alt) apply (erule order_less_le_trans) apply (simp add: unat_of_nat) done lemma nat_uint_less_helper: "nat (uint y) = z \ x < y \ nat (uint x) < z" apply (erule subst) apply (subst unat_def [symmetric]) apply (subst unat_def [symmetric]) by (simp add: unat_mono) lemma of_nat_0: "\of_nat n = (0::'a::len word); n < 2 ^ LENGTH('a)\ \ n = 0" by (drule unat_of_nat_eq, simp) lemma word_leq_le_minus_one: "\ x \ y; x \ 0 \ \ x - 1 < (y :: 'a :: len word)" apply (simp add: word_less_nat_alt word_le_nat_alt) apply (subst unat_minus_one) apply assumption apply (cases "unat x") apply (simp add: unat_eq_zero) apply arith done lemma of_nat_inj: "\x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\ \ (of_nat x = (of_nat y :: 'a :: len word)) = (x = y)" by (simp add: word_unat.norm_eq_iff [symmetric]) lemma map_prefixI: "prefix xs ys \ prefix (map f xs) (map f ys)" by (clarsimp simp: prefix_def) lemma if_Some_None_eq_None: "((if P then Some v else None) = None) = (\ P)" by simp lemma CollectPairFalse [iff]: "{(a,b). False} = {}" by (simp add: split_def) lemma if_conj_dist: "((if b then w else x) \ (if b then y else z) \ X) = ((if b then w \ y else x \ z) \ X)" by simp lemma if_P_True1: "Q \ (if P then True else Q)" by simp lemma if_P_True2: "Q \ (if P then Q else True)" by simp lemma list_all2_induct [consumes 1, case_names Nil Cons]: assumes lall: "list_all2 Q xs ys" and nilr: "P [] []" and consr: "\x xs y ys. \list_all2 Q xs ys; Q x y; P xs ys\ \ P (x # xs) (y # ys)" shows "P xs ys" using lall proof (induct rule: list_induct2 [OF list_all2_lengthD [OF lall]]) case 1 then show ?case by auto fact+ next case (2 x xs y ys) show ?case proof (rule consr) from "2.prems" show "list_all2 Q xs ys" and "Q x y" by simp_all then show "P xs ys" by (intro "2.hyps") qed qed lemma list_all2_induct_suffixeq [consumes 1, case_names Nil Cons]: assumes lall: "list_all2 Q as bs" and nilr: "P [] []" and consr: "\x xs y ys. \list_all2 Q xs ys; Q x y; P xs ys; suffix (x # xs) as; suffix (y # ys) bs\ \ P (x # xs) (y # ys)" shows "P as bs" proof - define as' where "as' == as" define bs' where "bs' == bs" have "suffix as as' \ suffix bs bs'" unfolding as'_def bs'_def by simp then show ?thesis using lall proof (induct rule: list_induct2 [OF list_all2_lengthD [OF lall]]) case 1 show ?case by fact next case (2 x xs y ys) show ?case proof (rule consr) from "2.prems" show "list_all2 Q xs ys" and "Q x y" by simp_all then show "P xs ys" using "2.hyps" "2.prems" by (auto dest: suffix_ConsD) from "2.prems" show "suffix (x # xs) as" and "suffix (y # ys) bs" by (auto simp: as'_def bs'_def) qed qed qed lemma upto_enum_step_shift: "\ is_aligned p n \ \ ([p , p + 2 ^ m .e. p + 2 ^ n - 1]) = map ((+) p) [0, 2 ^ m .e. 2 ^ n - 1]" apply (erule is_aligned_get_word_bits) prefer 2 apply (simp add: map_idI) apply (clarsimp simp: upto_enum_step_def) apply (frule is_aligned_no_overflow) apply (simp add: linorder_not_le [symmetric]) done lemma upto_enum_step_shift_red: "\ is_aligned p sz; sz < LENGTH('a); us \ sz \ \ [p :: 'a :: len word, p + 2 ^ us .e. p + 2 ^ sz - 1] = map (\x. p + of_nat x * 2 ^ us) [0 ..< 2 ^ (sz - us)]" apply (subst upto_enum_step_shift, assumption) apply (simp add: upto_enum_step_red) done lemma div_to_mult_word_lt: "\ (x :: 'a :: len word) \ y div z \ \ x * z \ y" apply (cases "z = 0") apply simp apply (simp add: word_neq_0_conv) apply (rule order_trans) apply (erule(1) word_mult_le_mono1) apply (simp add: unat_div) apply (rule order_le_less_trans [OF div_mult_le]) apply simp apply (rule word_div_mult_le) done lemma upto_enum_step_subset: "set [x, y .e. z] \ {x .. z}" apply (clarsimp simp: upto_enum_step_def linorder_not_less) apply (drule div_to_mult_word_lt) apply (rule conjI) apply (erule word_random[rotated]) apply simp apply (rule order_trans) apply (erule word_plus_mono_right) apply simp apply simp done lemma shiftr_less_t2n': "\ x && mask (n + m) = x; m < LENGTH('a) \ \ x >> n < 2 ^ m" for x :: "'a :: len word" apply (simp add: word_size mask_eq_iff_w2p[symmetric]) apply word_eqI apply (erule_tac x="na + n" in allE) apply fastforce done lemma shiftr_less_t2n: "x < 2 ^ (n + m) \ x >> n < 2 ^ m" for x :: "'a :: len word" apply (rule shiftr_less_t2n') apply (erule less_mask_eq) apply (rule ccontr) apply (simp add: not_less) apply (subst (asm) p2_eq_0[symmetric]) apply (simp add: power_add) done lemma shiftr_eq_0: "n \ LENGTH('a) \ ((w::'a::len word) >> n) = 0" apply (cut_tac shiftr_less_t2n'[of w n 0], simp) apply (simp add: mask_eq_iff) apply (simp add: lt2p_lem) apply simp done lemma shiftr_not_mask_0: "n+m \ LENGTH('a :: len) \ ((w::'a::len word) >> n) && ~~ (mask m) = 0" apply (simp add: and_not_mask shiftr_less_t2n shiftr_shiftr) apply (subgoal_tac "w >> n + m = 0", simp) apply (simp add: le_mask_iff[symmetric] mask_def le_def) apply (subst (asm) p2_gt_0[symmetric]) apply (simp add: power_add not_less) done lemma shiftl_less_t2n: fixes x :: "'a :: len word" shows "\ x < (2 ^ (m - n)); m < LENGTH('a) \ \ (x << n) < 2 ^ m" apply (simp add: word_size mask_eq_iff_w2p[symmetric]) apply word_eqI apply (erule_tac x="na - n" in allE) apply auto done lemma shiftl_less_t2n': "(x::'a::len word) < 2 ^ m \ m+n < LENGTH('a) \ x << n < 2 ^ (m + n)" by (rule shiftl_less_t2n) simp_all lemma ucast_ucast_mask: "(ucast :: 'a :: len word \ 'b :: len word) (ucast x) = x && mask (len_of TYPE ('a))" by word_eqI lemma ucast_ucast_len: "\ x < 2 ^ LENGTH('b) \ \ ucast (ucast x::'b::len word) = (x::'a::len word)" apply (subst ucast_ucast_mask) apply (erule less_mask_eq) done lemma ucast_ucast_id: "LENGTH('a) < LENGTH('b) \ ucast (ucast (x::'a::len word)::'b::len word) = x" by (auto intro: ucast_up_ucast_id simp: is_up_def source_size_def target_size_def word_size) lemma unat_ucast: - "unat (ucast x :: ('a :: len0) word) = unat x mod 2 ^ (LENGTH('a))" + "unat (ucast x :: ('a :: len) word) = unat x mod 2 ^ (LENGTH('a))" apply (simp add: unat_def ucast_def) apply (subst word_uint.eq_norm) apply (subst nat_mod_distrib) apply simp apply simp apply (subst nat_power_eq) apply simp apply simp done lemma ucast_less_ucast: "LENGTH('a) \ LENGTH('b) \ (ucast x < ((ucast (y :: 'a::len word)) :: 'b::len word)) = (x < y)" apply (simp add: word_less_nat_alt unat_ucast) apply (subst mod_less) apply(rule less_le_trans[OF unat_lt2p], simp) apply (subst mod_less) apply(rule less_le_trans[OF unat_lt2p], simp) apply simp done \ \This weaker version was previously called @{text ucast_less_ucast}. We retain it to support existing proofs.\ lemmas ucast_less_ucast_weak = ucast_less_ucast[OF order.strict_implies_order] lemma sints_subset: "m \ n \ sints m \ sints n" apply (simp add: sints_num) apply clarsimp apply (rule conjI) apply (erule order_trans[rotated]) apply simp apply (erule order_less_le_trans) apply simp done lemma up_scast_inj: "\ scast x = (scast y :: 'b :: len word); size x \ LENGTH('b) \ \ x = y" apply (simp add: scast_def) apply (subst(asm) word_sint.Abs_inject) apply (erule subsetD [OF sints_subset]) apply (simp add: word_size) apply (erule subsetD [OF sints_subset]) apply (simp add: word_size) apply simp done lemma up_scast_inj_eq: "LENGTH('a) \ len_of TYPE ('b) \ (scast x = (scast y::'b::len word)) = (x = (y::'a::len word))" by (fastforce dest: up_scast_inj simp: word_size) lemma nth_bounded: "\(x :: 'a :: len word) !! n; x < 2 ^ m; m \ len_of TYPE ('a)\ \ n < m" apply (frule test_bit_size) apply (clarsimp simp: test_bit_bl word_size) apply (simp add: nth_rev) apply (subst(asm) is_aligned_add_conv[OF is_aligned_0', simplified add_0_left, rotated]) apply assumption+ apply (simp only: to_bl_0) apply (simp add: nth_append split: if_split_asm) done lemma is_aligned_add_or: "\is_aligned p n; d < 2 ^ n\ \ p + d = p || d" by (rule word_plus_and_or_coroll, word_eqI) blast lemma two_power_increasing: "\ n \ m; m < LENGTH('a) \ \ (2 :: 'a :: len word) ^ n \ 2 ^ m" by (simp add: word_le_nat_alt) lemma is_aligned_add_less_t2n: "\is_aligned (p::'a::len word) n; d < 2^n; n \ m; p < 2^m\ \ p + d < 2^m" apply (case_tac "m < LENGTH('a)") apply (subst mask_eq_iff_w2p[symmetric]) apply (simp add: word_size) apply (simp add: is_aligned_add_or word_ao_dist less_mask_eq) apply (subst less_mask_eq) apply (erule order_less_le_trans) apply (erule(1) two_power_increasing) apply simp apply (simp add: power_overflow) done lemma aligned_offset_non_zero: "\ is_aligned x n; y < 2 ^ n; x \ 0 \ \ x + y \ 0" apply (cases "y = 0") apply simp apply (subst word_neq_0_conv) apply (subst gt0_iff_gem1) apply (erule is_aligned_get_word_bits) apply (subst field_simps[symmetric], subst plus_le_left_cancel_nowrap) apply (rule is_aligned_no_wrap') apply simp apply (rule word_leq_le_minus_one) apply simp apply assumption apply (erule (1) is_aligned_no_wrap') apply (simp add: gt0_iff_gem1 [symmetric] word_neq_0_conv) apply simp done lemmas mask_inner_mask = mask_eqs(1) lemma mask_add_aligned: "is_aligned p n \ (p + q) && mask n = q && mask n" apply (simp add: is_aligned_mask) apply (subst mask_inner_mask [symmetric]) apply simp done lemma take_prefix: "(take (length xs) ys = xs) = prefix xs ys" proof (induct xs arbitrary: ys) case Nil then show ?case by simp next case Cons then show ?case by (cases ys) auto qed lemma cart_singleton_empty: "(S \ {e} = {}) = (S = {})" by blast lemma word_div_1: "(n :: 'a :: len word) div 1 = n" by (simp add: word_div_def) lemma word_minus_one_le: "-1 \ (x :: 'a :: len word) = (x = -1)" apply (insert word_n1_ge[where y=x]) apply safe apply (erule(1) order_antisym) done lemma mask_out_sub_mask: "(x && ~~ (mask n)) = x - (x && (mask n))" by (simp add: field_simps word_plus_and_or_coroll2) lemma is_aligned_addD1: assumes al1: "is_aligned (x + y) n" and al2: "is_aligned (x::'a::len word) n" shows "is_aligned y n" using al2 proof (rule is_aligned_get_word_bits) assume "x = 0" then show ?thesis using al1 by simp next assume nv: "n < LENGTH('a)" from al1 obtain q1 where xy: "x + y = 2 ^ n * of_nat q1" and "q1 < 2 ^ (LENGTH('a) - n)" by (rule is_alignedE) moreover from al2 obtain q2 where x: "x = 2 ^ n * of_nat q2" and "q2 < 2 ^ (LENGTH('a) - n)" by (rule is_alignedE) ultimately have "y = 2 ^ n * (of_nat q1 - of_nat q2)" by (simp add: field_simps) then show ?thesis using nv by (simp add: is_aligned_mult_triv1) qed lemmas is_aligned_addD2 = is_aligned_addD1[OF subst[OF add.commute, of "%x. is_aligned x n" for n]] lemma is_aligned_add: "\is_aligned p n; is_aligned q n\ \ is_aligned (p + q) n" by (simp add: is_aligned_mask mask_add_aligned) lemma word_le_add: fixes x :: "'a :: len word" shows "x \ y \ \n. y = x + of_nat n" by (rule exI [where x = "unat (y - x)"]) simp lemma word_plus_mcs_4': fixes x :: "'a :: len word" shows "\x + v \ x + w; x \ x + v\ \ v \ w" apply (rule word_plus_mcs_4) apply (simp add: add.commute) apply (simp add: add.commute) done lemma shiftl_mask_is_0[simp]: "(x << n) && mask n = 0" apply (rule iffD1 [OF is_aligned_mask]) apply (rule is_aligned_shiftl_self) done definition sum_map :: "('a \ 'b) \ ('c \ 'd) \ 'a + 'c \ 'b + 'd" where "sum_map f g x \ case x of Inl v \ Inl (f v) | Inr v' \ Inr (g v')" lemma sum_map_simps[simp]: "sum_map f g (Inl v) = Inl (f v)" "sum_map f g (Inr w) = Inr (g w)" by (simp add: sum_map_def)+ lemma if_and_helper: "(If x v v') && v'' = If x (v && v'') (v' && v'')" by (rule if_distrib) lemma unat_Suc2: fixes n :: "'a :: len word" shows "n \ -1 \ unat (n + 1) = Suc (unat n)" apply (subst add.commute, rule unatSuc) apply (subst eq_diff_eq[symmetric], simp add: minus_equation_iff) done lemmas word_unat_Rep_inject1 = word_unat.Rep_inject[where y=1] lemmas unat_eq_1 = unat_eq_0 word_unat_Rep_inject1[simplified] lemma rshift_sub_mask_eq: "(a >> (size a - b)) && mask b = a >> (size a - b)" using shiftl_shiftr2[where a=a and b=0 and c="size a - b"] apply (cases "b < size a") apply simp apply (simp add: linorder_not_less mask_def word_size p2_eq_0[THEN iffD2]) done lemma shiftl_shiftr3: "b \ c \ a << b >> c = (a >> c - b) && mask (size a - c)" apply (cases "b = c") apply (simp add: shiftl_shiftr1) apply (simp add: shiftl_shiftr2) done lemma and_mask_shiftr_comm: "m\size w \ (w && mask m) >> n = (w >> n) && mask (m-n)" by (simp add: and_mask shiftr_shiftr) (simp add: word_size shiftl_shiftr3) lemma and_mask_shiftl_comm: "m+n \ size w \ (w && mask m) << n = (w << n) && mask (m+n)" by (simp add: and_mask word_size shiftl_shiftl) (simp add: shiftl_shiftr1) lemma le_mask_shiftl_le_mask: "s = m + n \ x \ mask n \ x << m \ mask s" by (simp add: le_mask_iff shiftl_shiftr3) lemma and_not_mask_twice: "(w && ~~ (mask n)) && ~~ (mask m) = w && ~~ (mask (max m n))" apply (simp add: and_not_mask) apply (case_tac "n x = y - 1 \ x < y - (1 ::'a::len word)" apply (drule word_less_sub_1) apply (drule order_le_imp_less_or_eq) apply auto done lemma eq_eqI: "a = b \ (a = x) = (b = x)" by simp lemma mask_and_mask: "mask a && mask b = mask (min a b)" by word_eqI lemma mask_eq_0_eq_x: "(x && w = 0) = (x && ~~ w = x)" using word_plus_and_or_coroll2[where x=x and w=w] by auto lemma mask_eq_x_eq_0: "(x && w = x) = (x && ~~ w = 0)" using word_plus_and_or_coroll2[where x=x and w=w] by auto definition "limited_and (x :: 'a :: len word) y = (x && y = x)" lemma limited_and_eq_0: "\ limited_and x z; y && ~~ z = y \ \ x && y = 0" unfolding limited_and_def apply (subst arg_cong2[where f="(&&)"]) apply (erule sym)+ apply (simp(no_asm) add: word_bw_assocs word_bw_comms word_bw_lcs) done lemma limited_and_eq_id: "\ limited_and x z; y && z = z \ \ x && y = x" unfolding limited_and_def by (erule subst, fastforce simp: word_bw_lcs word_bw_assocs word_bw_comms) lemma lshift_limited_and: "limited_and x z \ limited_and (x << n) (z << n)" unfolding limited_and_def by (simp add: shiftl_over_and_dist[symmetric]) lemma rshift_limited_and: "limited_and x z \ limited_and (x >> n) (z >> n)" unfolding limited_and_def by (simp add: shiftr_over_and_dist[symmetric]) lemmas limited_and_simps1 = limited_and_eq_0 limited_and_eq_id lemmas is_aligned_limited_and = is_aligned_neg_mask_eq[unfolded mask_def, folded limited_and_def] lemma compl_of_1: "~~ 1 = (-2 :: 'a :: len word)" by simp lemmas limited_and_simps = limited_and_simps1 limited_and_simps1[OF is_aligned_limited_and] limited_and_simps1[OF lshift_limited_and] limited_and_simps1[OF rshift_limited_and] limited_and_simps1[OF rshift_limited_and, OF is_aligned_limited_and] compl_of_1 shiftl_shiftr1[unfolded word_size mask_def] shiftl_shiftr2[unfolded word_size mask_def] lemma split_word_eq_on_mask: "(x = y) = (x && m = y && m \ x && ~~ m = y && ~~ m)" by safe word_eqI_solve lemma map2_Cons_2_3: "(map2 f xs (y # ys) = (z # zs)) = (\x xs'. xs = x # xs' \ f x y = z \ map2 f xs' ys = zs)" by (case_tac xs, simp_all) lemma map2_xor_replicate_False: "map2 (\x y. x \ \ y) xs (replicate n False) = take n xs" apply (induct xs arbitrary: n, simp) apply (case_tac n; simp) done lemma word_and_1_shiftl: "x && (1 << n) = (if x !! n then (1 << n) else 0)" for x :: "'a :: len word" by word_eqI_solve lemmas word_and_1_shiftls' = word_and_1_shiftl[where n=0] word_and_1_shiftl[where n=1] word_and_1_shiftl[where n=2] lemmas word_and_1_shiftls = word_and_1_shiftls' [simplified] lemma word_and_mask_shiftl: "x && (mask n << m) = ((x >> m) && mask n) << m" by word_eqI_solve lemma plus_Collect_helper: "(+) x ` {xa. P (xa :: 'a :: len word)} = {xa. P (xa - x)}" by (fastforce simp add: image_def) lemma plus_Collect_helper2: "(+) (- x) ` {xa. P (xa :: 'a :: len word)} = {xa. P (x + xa)}" using plus_Collect_helper [of "- x" P] by (simp add: ac_simps) lemma word_FF_is_mask: "0xFF = mask 8" by (simp add: mask_def) lemma word_1FF_is_mask: "0x1FF = mask 9" by (simp add: mask_def) lemma ucast_of_nat_small: "x < 2 ^ LENGTH('a) \ ucast (of_nat x :: 'a :: len word) = (of_nat x :: 'b :: len word)" apply (rule sym, subst word_unat.inverse_norm) apply (simp add: ucast_def word_of_int[symmetric] of_nat_nat[symmetric] unat_def[symmetric]) apply (simp add: unat_of_nat) done lemma word_le_make_less: fixes x :: "'a :: len word" shows "y \ -1 \ (x \ y) = (x < (y + 1))" apply safe apply (erule plus_one_helper2) apply (simp add: eq_diff_eq[symmetric]) done lemmas finite_word = finite [where 'a="'a::len word"] lemma word_to_1_set: "{0 ..< (1 :: 'a :: len word)} = {0}" by fastforce lemma range_subset_eq2: "{a :: 'a :: len word .. b} \ {} \ ({a .. b} \ {c .. d}) = (c \ a \ b \ d)" by simp lemma word_leq_minus_one_le: fixes x :: "'a::len word" shows "\y \ 0; x \ y - 1 \ \ x < y" using le_m1_iff_lt word_neq_0_conv by blast lemma word_count_from_top: "n \ 0 \ {0 ..< n :: 'a :: len word} = {0 ..< n - 1} \ {n - 1}" apply (rule set_eqI, rule iffI) apply simp apply (drule word_le_minus_one_leq) apply (rule disjCI) apply simp apply simp apply (erule word_leq_minus_one_le) apply fastforce done lemma word_minus_one_le_leq: "\ x - 1 < y \ \ x \ (y :: 'a :: len word)" apply (cases "x = 0") apply simp apply (simp add: word_less_nat_alt word_le_nat_alt) apply (subst(asm) unat_minus_one) apply (simp add: word_less_nat_alt) apply (cases "unat x") apply (simp add: unat_eq_zero) apply arith done lemma mod_mod_power: fixes k :: nat shows "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ (min m n)" proof (cases "m \ n") case True then have "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ m" apply - apply (subst mod_less [where n = "2 ^ n"]) apply (rule order_less_le_trans [OF mod_less_divisor]) apply simp+ done also have "\ = k mod 2 ^ (min m n)" using True by simp finally show ?thesis . next case False then have "n < m" by simp then obtain d where md: "m = n + d" by (auto dest: less_imp_add_positive) then have "k mod 2 ^ m = 2 ^ n * (k div 2 ^ n mod 2 ^ d) + k mod 2 ^ n" by (simp add: mod_mult2_eq power_add) then have "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ n" by (simp add: mod_add_left_eq) then show ?thesis using False by simp qed lemma word_div_less: "m < n \ m div n = 0" for m :: "'a :: len word" by (simp add: unat_mono word_arith_nat_defs(6)) lemma word_must_wrap: "\ x \ n - 1; n \ x \ \ n = (0 :: 'a :: len word)" using dual_order.trans sub_wrap word_less_1 by blast lemma range_subset_card: "\ {a :: 'a :: len word .. b} \ {c .. d}; b \ a \ \ d \ c \ d - c \ b - a" using word_sub_le word_sub_mono by fastforce lemma less_1_simp: "n - 1 < m = (n \ (m :: 'a :: len word) \ n \ 0)" by unat_arith lemma alignUp_div_helper: fixes a :: "'a::len word" assumes kv: "k < 2 ^ (LENGTH('a) - n)" and xk: "x = 2 ^ n * of_nat k" and le: "a \ x" and sz: "n < LENGTH('a)" and anz: "a mod 2 ^ n \ 0" shows "a div 2 ^ n < of_nat k" proof - have kn: "unat (of_nat k :: 'a word) * unat ((2::'a word) ^ n) < 2 ^ LENGTH('a)" using xk kv sz apply (subst unat_of_nat_eq) apply (erule order_less_le_trans) apply simp apply (subst unat_power_lower, simp) apply (subst mult.commute) apply (rule nat_less_power_trans) apply simp apply simp done have "unat a div 2 ^ n * 2 ^ n \ unat a" proof - have "unat a = unat a div 2 ^ n * 2 ^ n + unat a mod 2 ^ n" by (simp add: div_mult_mod_eq) also have "\ \ unat a div 2 ^ n * 2 ^ n" using sz anz by (simp add: unat_arith_simps) finally show ?thesis .. qed then have "a div 2 ^ n * 2 ^ n < a" using sz anz apply (subst word_less_nat_alt) apply (subst unat_word_ariths) apply (subst unat_div) apply simp apply (rule order_le_less_trans [OF mod_less_eq_dividend]) apply (erule order_le_neq_trans [OF div_mult_le]) done also from xk le have "\ \ of_nat k * 2 ^ n" by (simp add: field_simps) finally show ?thesis using sz kv apply - apply (erule word_mult_less_dest [OF _ _ kn]) apply (simp add: unat_div) apply (rule order_le_less_trans [OF div_mult_le]) apply (rule unat_lt2p) done qed lemma nat_mod_power_lem: fixes a :: nat shows "1 < a \ a ^ n mod a ^ m = (if m \ n then 0 else a ^ n)" apply (clarsimp) apply (clarsimp simp add: le_iff_add power_add) done lemma power_mod_div: fixes x :: "nat" shows "x mod 2 ^ n div 2 ^ m = x div 2 ^ m mod 2 ^ (n - m)" (is "?LHS = ?RHS") proof (cases "n \ m") case True then have "?LHS = 0" apply - apply (rule div_less) apply (rule order_less_le_trans [OF mod_less_divisor]; simp) done also have "\ = ?RHS" using True by simp finally show ?thesis . next case False then have lt: "m < n" by simp then obtain q where nv: "n = m + q" and "0 < q" by (auto dest: less_imp_Suc_add) then have "x mod 2 ^ n = 2 ^ m * (x div 2 ^ m mod 2 ^ q) + x mod 2 ^ m" by (simp add: power_add mod_mult2_eq) then have "?LHS = x div 2 ^ m mod 2 ^ q" by (simp add: div_add1_eq) also have "\ = ?RHS" using nv by simp finally show ?thesis . qed lemma word_power_mod_div: fixes x :: "'a::len word" shows "\ n < LENGTH('a); m < LENGTH('a)\ \ x mod 2 ^ n div 2 ^ m = x div 2 ^ m mod 2 ^ (n - m)" apply (simp add: word_arith_nat_div unat_mod power_mod_div) apply (subst unat_arith_simps(3)) apply (subst unat_mod) apply (subst unat_of_nat)+ apply (simp add: mod_mod_power min.commute) done lemma word_range_minus_1': fixes a :: "'a :: len word" shows "a \ 0 \ {a - 1<..b} = {a..b}" by (simp add: greaterThanAtMost_def atLeastAtMost_def greaterThan_def atLeast_def less_1_simp) lemma word_range_minus_1: fixes a :: "'a :: len word" shows "b \ 0 \ {a..b - 1} = {a.. 'b :: len word) x" by (simp add: ucast_def word_of_int_nat unat_def) lemmas if_fun_split = if_apply_def2 lemma i_hate_words_helper: "i \ (j - k :: nat) \ i \ j" by simp lemma i_hate_words: "unat (a :: 'a word) \ unat (b :: 'a :: len word) - Suc 0 \ a \ -1" apply (frule i_hate_words_helper) apply (subst(asm) word_le_nat_alt[symmetric]) apply (clarsimp simp only: word_minus_one_le) apply (simp only: linorder_not_less[symmetric]) apply (erule notE) apply (rule diff_Suc_less) apply (subst neq0_conv[symmetric]) apply (subst unat_eq_0) apply (rule notI, drule arg_cong[where f="(+) 1"]) apply simp done lemma overflow_plus_one_self: "(1 + p \ p) = (p = (-1 :: 'a :: len word))" apply rule apply (rule ccontr) apply (drule plus_one_helper2) apply (rule notI) apply (drule arg_cong[where f="\x. x - 1"]) apply simp apply (simp add: field_simps) apply simp done lemma plus_1_less: "(x + 1 \ (x :: 'a :: len word)) = (x = -1)" apply (rule iffI) apply (rule ccontr) apply (cut_tac plus_one_helper2[where x=x, OF order_refl]) apply simp apply clarsimp apply (drule arg_cong[where f="\x. x - 1"]) apply simp apply simp done lemma pos_mult_pos_ge: "[|x > (0::int); n>=0 |] ==> n * x >= n*1" apply (simp only: mult_left_mono) done lemma If_eq_obvious: "x \ z \ ((if P then x else y) = z) = (\ P \ y = z)" by simp lemma Some_to_the: "v = Some x \ x = the v" by simp lemma dom_if_Some: "dom (\x. if P x then Some (f x) else g x) = {x. P x} \ dom g" by fastforce lemma dom_insert_absorb: "x \ dom f \ insert x (dom f) = dom f" by auto lemma emptyE2: "\ S = {}; x \ S \ \ P" by simp lemma mod_div_equality_div_eq: "a div b * b = (a - (a mod b) :: int)" by (simp add: field_simps) lemma zmod_helper: "n mod m = k \ ((n :: int) + a) mod m = (k + a) mod m" by (metis add.commute mod_add_right_eq) lemma int_div_sub_1: "\ m \ 1 \ \ (n - (1 :: int)) div m = (if m dvd n then (n div m) - 1 else n div m)" apply (subgoal_tac "m = 0 \ (n - (1 :: int)) div m = (if m dvd n then (n div m) - 1 else n div m)") apply fastforce apply (subst mult_cancel_right[symmetric]) apply (simp only: left_diff_distrib split: if_split) apply (simp only: mod_div_equality_div_eq) apply (clarsimp simp: field_simps) apply (clarsimp simp: dvd_eq_mod_eq_0) apply (cases "m = 1") apply simp apply (subst mod_diff_eq[symmetric], simp add: zmod_minus1) apply clarsimp apply (subst diff_add_cancel[where b=1, symmetric]) apply (subst mod_add_eq[symmetric]) apply (simp add: field_simps) apply (rule mod_pos_pos_trivial) apply (subst add_0_right[where a=0, symmetric]) apply (rule add_mono) apply simp apply simp apply (cases "(n - 1) mod m = m - 1") apply (drule zmod_helper[where a=1]) apply simp apply (subgoal_tac "1 + (n - 1) mod m \ m") apply simp apply (subst field_simps, rule zless_imp_add1_zle) apply simp done lemma ptr_add_image_multI: "\ \x y. (x * val = y * val') = (x * val'' = y); x * val'' \ S \ \ ptr_add ptr (x * val) \ (\p. ptr_add ptr (p * val')) ` S" apply (simp add: image_def) apply (erule rev_bexI) apply (rule arg_cong[where f="ptr_add ptr"]) apply simp done lemma shift_times_fold: "(x :: 'a :: len word) * (2 ^ n) << m = x << (m + n)" by (simp add: shiftl_t2n ac_simps power_add) lemma word_plus_strict_mono_right: fixes x :: "'a :: len word" shows "\y < z; x \ x + z\ \ x + y < x + z" by unat_arith lemma replicate_minus: "k < n \ replicate n False = replicate (n - k) False @ replicate k False" by (subst replicate_add [symmetric]) simp lemmas map_prod_split_imageI' = map_prod_imageI[where f="case_prod f" and g="case_prod g" and a="(a, b)" and b="(c, d)" for a b c d f g] lemmas map_prod_split_imageI = map_prod_split_imageI'[simplified] lemma word_div_mult: "0 < c \ a < b * c \ a div c < b" for a b c :: "'a::len word" by (rule classical) (use div_to_mult_word_lt [of b a c] in \auto simp add: word_less_nat_alt word_le_nat_alt unat_div\) lemma word_less_power_trans_ofnat: "\n < 2 ^ (m - k); k \ m; m < LENGTH('a)\ \ of_nat n * 2 ^ k < (2::'a::len word) ^ m" apply (subst mult.commute) apply (rule word_less_power_trans) apply (simp add: word_less_nat_alt) apply (subst unat_of_nat_eq) apply (erule order_less_trans) apply simp+ done lemma word_1_le_power: "n < LENGTH('a) \ (1 :: 'a :: len word) \ 2 ^ n" by (rule inc_le[where i=0, simplified], erule iffD2[OF p2_gt_0]) lemma enum_word_div: fixes v :: "'a :: len word" shows "\xs ys. enum = xs @ [v] @ ys \ (\x \ set xs. x < v) \ (\y \ set ys. v < y)" apply (simp only: enum_word_def) apply (subst upt_add_eq_append'[where j="unat v"]) apply simp apply (rule order_less_imp_le, simp) apply (simp add: upt_conv_Cons) apply (intro exI conjI) apply fastforce apply clarsimp apply (drule of_nat_mono_maybe[rotated, where 'a='a]) apply simp apply simp apply (clarsimp simp: Suc_le_eq) apply (drule of_nat_mono_maybe[rotated, where 'a='a]) apply simp apply simp done lemma of_bool_nth: "of_bool (x !! v) = (x >> v) && 1" by (word_eqI cong: rev_conj_cong) lemma unat_1_0: "1 \ (x::'a::len word) = (0 < unat x)" by (auto simp add: word_le_nat_alt) lemma x_less_2_0_1': fixes x :: "'a::len word" shows "\LENGTH('a) \ 1; x < 2\ \ x = 0 \ x = 1" apply (induct x) apply clarsimp+ by (metis Suc_eq_plus1 add_lessD1 less_irrefl one_add_one unatSuc word_less_nat_alt) lemmas word_add_le_iff2 = word_add_le_iff [folded no_olen_add_nat] lemma of_nat_power: shows "\ p < 2 ^ x; x < len_of TYPE ('a) \ \ of_nat p < (2 :: 'a :: len word) ^ x" apply (rule order_less_le_trans) apply (rule of_nat_mono_maybe) apply (erule power_strict_increasing) apply simp apply assumption apply (simp add: word_unat_power) done lemma of_nat_n_less_equal_power_2: "n < LENGTH('a::len) \ ((of_nat n)::'a word) < 2 ^ n" apply (induct n) apply clarsimp apply clarsimp apply (metis of_nat_power n_less_equal_power_2 of_nat_Suc power_Suc) done lemma eq_mask_less: fixes w :: "'a::len word" assumes eqm: "w = w && mask n" and sz: "n < len_of TYPE ('a)" shows "w < (2::'a word) ^ n" by (subst eqm, rule and_mask_less' [OF sz]) lemma of_nat_mono_maybe': fixes Y :: "nat" assumes xlt: "x < 2 ^ len_of TYPE ('a)" assumes ylt: "y < 2 ^ len_of TYPE ('a)" shows "(y < x) = (of_nat y < (of_nat x :: 'a :: len word))" apply (subst word_less_nat_alt) apply (subst unat_of_nat)+ apply (subst mod_less) apply (rule ylt) apply (subst mod_less) apply (rule xlt) apply simp done lemma shiftr_mask_eq: "(x >> n) && mask (size x - n) = x >> n" for x :: "'a :: len word" by word_eqI_solve lemma shiftr_mask_eq': "m = (size x - n) \ (x >> n) && mask m = x >> n" for x :: "'a :: len word" by (simp add: shiftr_mask_eq) lemma dom_if: "dom (\a. if a \ addrs then Some (f a) else g a) = addrs \ dom g" by (auto simp: dom_def split: if_split) lemma less_is_non_zero_p1: fixes a :: "'a :: len word" shows "a < k \ a + 1 \ 0" apply (erule contrapos_pn) apply (drule max_word_wrap) apply (simp add: not_less) done lemma of_nat_mono_maybe_le: "\x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\ \ (y \ x) = ((of_nat y :: 'a :: len word) \ of_nat x)" apply (clarsimp simp: le_less) apply (rule disj_cong) apply (rule of_nat_mono_maybe', assumption+) apply (simp add: word_unat.norm_eq_iff [symmetric]) done lemma mask_AND_NOT_mask: "(w && ~~ (mask n)) && mask n = 0" by word_eqI lemma AND_NOT_mask_plus_AND_mask_eq: "(w && ~~ (mask n)) + (w && mask n) = w" by (subst word_plus_and_or_coroll; word_eqI_solve) lemma mask_eqI: fixes x :: "'a :: len word" assumes m1: "x && mask n = y && mask n" and m2: "x && ~~ (mask n) = y && ~~ (mask n)" shows "x = y" proof (subst bang_eq, rule allI) fix m show "x !! m = y !! m" proof (cases "m < n") case True then have "x !! m = ((x && mask n) !! m)" by (simp add: word_size test_bit_conj_lt) also have "\ = ((y && mask n) !! m)" using m1 by simp also have "\ = y !! m" using True by (simp add: word_size test_bit_conj_lt) finally show ?thesis . next case False then have "x !! m = ((x && ~~ (mask n)) !! m)" by (simp add: neg_mask_test_bit test_bit_conj_lt) also have "\ = ((y && ~~ (mask n)) !! m)" using m2 by simp also have "\ = y !! m" using False by (simp add: neg_mask_test_bit test_bit_conj_lt) finally show ?thesis . qed qed lemma nat_less_power_trans2: fixes n :: nat shows "\n < 2 ^ (m - k); k \ m\ \ n * 2 ^ k < 2 ^ m" by (subst mult.commute, erule (1) nat_less_power_trans) lemma nat_move_sub_le: "(a::nat) + b \ c \ a \ c - b" by arith lemma neq_0_no_wrap: fixes x :: "'a :: len word" shows "\ x \ x + y; x \ 0 \ \ x + y \ 0" by clarsimp lemma plus_minus_one_rewrite: "v + (- 1 :: ('a :: {ring, one, uminus})) \ v - 1" by (simp add: field_simps) lemma power_minus_is_div: "b \ a \ (2 :: nat) ^ (a - b) = 2 ^ a div 2 ^ b" apply (induct a arbitrary: b) apply simp apply (erule le_SucE) apply (clarsimp simp:Suc_diff_le le_iff_add power_add) apply simp done lemma two_pow_div_gt_le: "v < 2 ^ n div (2 ^ m :: nat) \ m \ n" by (clarsimp dest!: less_two_pow_divD) lemma unatSuc2: fixes n :: "'a :: len word" shows "n + 1 \ 0 \ unat (n + 1) = Suc (unat n)" by (simp add: add.commute unatSuc) lemma word_of_nat_less: "\ n < unat x \ \ of_nat n < x" apply (simp add: word_less_nat_alt) apply (erule order_le_less_trans[rotated]) apply (simp add: unat_of_nat) done lemma word_of_nat_le: "n \ unat x \ of_nat n \ x" apply (simp add: word_le_nat_alt unat_of_nat) apply (erule order_trans[rotated]) apply simp done lemma word_unat_less_le: "a \ of_nat b \ unat a \ b" by (metis eq_iff le_cases le_unat_uoi word_of_nat_le) lemma and_eq_0_is_nth: fixes x :: "'a :: len word" shows "y = 1 << n \ ((x && y) = 0) = (\ (x !! n))" apply safe apply (drule_tac u="(x && (1 << n))" and x=n in word_eqD) apply (simp add: nth_w2p) apply (simp add: test_bit_bin) apply word_eqI done lemmas arg_cong_Not = arg_cong [where f=Not] lemmas and_neq_0_is_nth = arg_cong_Not [OF and_eq_0_is_nth, simplified] lemma nth_is_and_neq_0: "(x::'a::len word) !! n = (x && 2 ^ n \ 0)" by (subst and_neq_0_is_nth; rule refl) lemma mask_Suc_0 : "mask (Suc 0) = 1" by (simp add: mask_def) lemma ucast_ucast_add: fixes x :: "'a :: len word" fixes y :: "'b :: len word" shows "LENGTH('b) \ LENGTH('a) \ ucast (ucast x + y) = x + ucast y" apply (rule word_unat.Rep_eqD) apply (simp add: unat_ucast unat_word_ariths mod_mod_power min.absorb2 unat_of_nat) apply (subst mod_add_left_eq[symmetric]) apply (simp add: mod_mod_power min.absorb2) apply (subst mod_add_right_eq) apply simp done lemma word_shift_zero: "\ x << n = 0; x \ 2^m; m + n < LENGTH('a)\ \ (x::'a::len word) = 0" apply (rule ccontr) apply (drule (2) word_shift_nonzero) apply simp done lemma bool_mask': fixes x :: "'a :: len word" shows "2 < LENGTH('a) \ (0 < x && 1) = (x && 1 = 1)" apply (rule iffI) prefer 2 apply simp apply (subgoal_tac "x && mask 1 < 2^1") prefer 2 apply (rule and_mask_less_size) apply (simp add: word_size) apply (simp add: mask_def) apply (drule word_less_cases [where y=2]) apply (erule disjE, simp) apply simp done lemma sint_eq_uint: "\ msb x \ sint x = uint x" apply (rule word_uint.Abs_eqD, subst word_sint.Rep_inverse) apply simp_all apply (cut_tac x=x in word_sint.Rep) apply (clarsimp simp add: uints_num sints_num) apply (rule conjI) apply (rule ccontr) apply (simp add: linorder_not_le word_msb_sint[symmetric]) apply (erule order_less_le_trans) apply simp done lemma scast_eq_ucast: "\ msb x \ scast x = ucast x" by (simp add: scast_def ucast_def sint_eq_uint) lemma lt1_neq0: fixes x :: "'a :: len word" shows "(1 \ x) = (x \ 0)" by unat_arith lemma word_plus_one_nonzero: fixes x :: "'a :: len word" shows "\x \ x + y; y \ 0\ \ x + 1 \ 0" apply (subst lt1_neq0 [symmetric]) apply (subst olen_add_eqv [symmetric]) apply (erule word_random) apply (simp add: lt1_neq0) done lemma word_sub_plus_one_nonzero: fixes n :: "'a :: len word" shows "\n' \ n; n' \ 0\ \ (n - n') + 1 \ 0" apply (subst lt1_neq0 [symmetric]) apply (subst olen_add_eqv [symmetric]) apply (rule word_random [where x' = n']) apply simp apply (erule word_sub_le) apply (simp add: lt1_neq0) done lemma word_le_minus_mono_right: fixes x :: "'a :: len word" shows "\ z \ y; y \ x; z \ x \ \ x - y \ x - z" apply (rule word_sub_mono) apply simp apply assumption apply (erule word_sub_le) apply (erule word_sub_le) done lemma drop_append_miracle: "n = length xs \ drop n (xs @ ys) = ys" by simp lemma foldr_does_nothing_to_xf: "\ \x s. x \ set xs \ xf (f x s) = xf s \ \ xf (foldr f xs s) = xf s" by (induct xs, simp_all) lemma nat_less_mult_monoish: "\ a < b; c < (d :: nat) \ \ (a + 1) * (c + 1) <= b * d" apply (drule Suc_leI)+ apply (drule(1) mult_le_mono) apply simp done lemma word_0_sle_from_less[unfolded word_size]: "\ x < 2 ^ (size x - 1) \ \ 0 <=s x" apply (clarsimp simp: word_sle_msb_le) apply (simp add: word_msb_nth) apply (subst (asm) word_test_bit_def [symmetric]) apply (drule less_mask_eq) apply (drule_tac x="size x - 1" in word_eqD) apply (simp add: word_size) done lemma not_msb_from_less: "(v :: 'a word) < 2 ^ (LENGTH('a :: len) - 1) \ \ msb v" apply (clarsimp simp add: msb_nth) apply (drule less_mask_eq) apply (drule word_eqD, drule(1) iffD2) apply simp done lemma distinct_lemma: "f x \ f y \ x \ y" by auto lemma ucast_sub_ucast: fixes x :: "'a::len word" assumes "y \ x" assumes T: "LENGTH('a) \ LENGTH('b)" shows "ucast (x - y) = (ucast x - ucast y :: 'b::len word)" proof - from T have P: "unat x < 2 ^ LENGTH('b)" "unat y < 2 ^ LENGTH('b)" by (fastforce intro!: less_le_trans[OF unat_lt2p])+ then show ?thesis by (simp add: unat_arith_simps unat_ucast assms[simplified unat_arith_simps]) qed lemma word_1_0: "\a + (1::('a::len) word) \ b; a < of_nat x\ \ a < b" by unat_arith lemma unat_of_nat_less:"\ a < b; unat b = c \ \ a < of_nat c" by fastforce lemma word_le_plus_1: "\ (y::('a::len) word) < y + n; a < n \ \ y + a \ y + a + 1" by unat_arith lemma word_le_plus:"\(a::('a::len) word) < a + b; c < b\ \ a \ a + c" by (metis order_less_imp_le word_random) (* * Basic signed arithemetic properties. *) lemma sint_minus1 [simp]: "(sint x = -1) = (x = -1)" by (metis sint_n1 word_sint.Rep_inverse') lemma sint_0 [simp]: "(sint x = 0) = (x = 0)" by (metis sint_0 word_sint.Rep_inverse') (* It is not always that case that "sint 1 = 1", because of 1-bit word sizes. * This lemma produces the different cases. *) lemma sint_1_cases: "\ \ len_of TYPE ('a::len) = 1; (a::'a word) = 0; sint a = 0 \ \ P; \ len_of TYPE ('a) = 1; a = 1; sint (1 :: 'a word) = -1 \ \ P; \ len_of TYPE ('a) > 1; sint (1 :: 'a word) = 1 \ \ P \ \ P" apply atomize_elim apply (case_tac "len_of TYPE ('a) = 1") apply clarsimp apply (subgoal_tac "(UNIV :: 'a word set) = {0, 1}") apply (metis UNIV_I insert_iff singletonE) apply (subst word_unat.univ) apply (clarsimp simp: unats_def image_def) apply (rule set_eqI, rule iffI) apply clarsimp apply (metis One_nat_def less_2_cases of_nat_1 semiring_1_class.of_nat_0) apply clarsimp apply (metis Abs_fnat_hom_0 Suc_1 lessI of_nat_1 zero_less_Suc) apply clarsimp apply (metis One_nat_def arith_is_1 le_def len_gt_0) done lemma sint_int_min: "sint (- (2 ^ (LENGTH('a) - Suc 0)) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))" apply (subst word_sint.Abs_inverse' [where r="- (2 ^ (LENGTH('a) - Suc 0))"]) apply (clarsimp simp: sints_num) apply (clarsimp simp: wi_hom_syms word_of_int_2p) apply clarsimp done lemma sint_int_max_plus_1: "sint (2 ^ (LENGTH('a) - Suc 0) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))" apply (subst word_of_int_2p [symmetric]) apply (subst int_word_sint) apply clarsimp apply (metis Suc_pred int_word_uint len_gt_0 power_Suc uint_eq_0 word_of_int_2p word_pow_0) done lemma sbintrunc_eq_in_range: "(sbintrunc n x = x) = (x \ range (sbintrunc n))" "(x = sbintrunc n x) = (x \ range (sbintrunc n))" apply (simp_all add: image_def) apply (metis sbintrunc_sbintrunc)+ done lemma sbintrunc_If: "- 3 * (2 ^ n) \ x \ x < 3 * (2 ^ n) \ sbintrunc n x = (if x < - (2 ^ n) then x + 2 * (2 ^ n) else if x \ 2 ^ n then x - 2 * (2 ^ n) else x)" apply (simp add: no_sbintr_alt2, safe) apply (simp add: mod_pos_geq) apply (subst mod_add_self1[symmetric], simp) done lemma signed_arith_eq_checks_to_ord: "(sint a + sint b = sint (a + b )) = ((a <=s a + b) = (0 <=s b))" "(sint a - sint b = sint (a - b )) = ((0 <=s a - b) = (b <=s a))" "(- sint a = sint (- a)) = (0 <=s (- a) = (a <=s 0))" using sint_range'[where x=a] sint_range'[where x=b] by (simp_all add: sint_word_ariths word_sle_def word_sless_alt sbintrunc_If) (* Basic proofs that signed word div/mod operations are * truncations of their integer counterparts. *) lemma signed_div_arith: "sint ((a::('a::len) word) sdiv b) = sbintrunc (LENGTH('a) - 1) (sint a sdiv sint b)" apply (subst word_sbin.norm_Rep [symmetric]) apply (subst bin_sbin_eq_iff' [symmetric]) apply simp apply (subst uint_sint [symmetric]) apply (clarsimp simp: sdiv_int_def sdiv_word_def) apply (metis word_ubin.eq_norm) done lemma signed_mod_arith: "sint ((a::('a::len) word) smod b) = sbintrunc (LENGTH('a) - 1) (sint a smod sint b)" apply (subst word_sbin.norm_Rep [symmetric]) apply (subst bin_sbin_eq_iff' [symmetric]) apply simp apply (subst uint_sint [symmetric]) apply (clarsimp simp: smod_int_def smod_word_def) apply (metis word_ubin.eq_norm) done (* Signed word arithmetic overflow constraints. *) lemma signed_arith_ineq_checks_to_eq: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) = (sint a + sint b = sint (a + b ))" "((- (2 ^ (size a - 1)) \ (sint a - sint b)) \ (sint a - sint b \ (2 ^ (size a - 1) - 1))) = (sint a - sint b = sint (a - b))" "((- (2 ^ (size a - 1)) \ (- sint a)) \ (- sint a) \ (2 ^ (size a - 1) - 1)) = ((- sint a) = sint (- a))" "((- (2 ^ (size a - 1)) \ (sint a * sint b)) \ (sint a * sint b \ (2 ^ (size a - 1) - 1))) = (sint a * sint b = sint (a * b))" "((- (2 ^ (size a - 1)) \ (sint a sdiv sint b)) \ (sint a sdiv sint b \ (2 ^ (size a - 1) - 1))) = (sint a sdiv sint b = sint (a sdiv b))" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) = (sint a smod sint b = sint (a smod b))" by (auto simp: sint_word_ariths word_size signed_div_arith signed_mod_arith sbintrunc_eq_in_range range_sbintrunc) lemma signed_arith_sint: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) \ sint (a + b) = (sint a + sint b)" "((- (2 ^ (size a - 1)) \ (sint a - sint b)) \ (sint a - sint b \ (2 ^ (size a - 1) - 1))) \ sint (a - b) = (sint a - sint b)" "((- (2 ^ (size a - 1)) \ (- sint a)) \ (- sint a) \ (2 ^ (size a - 1) - 1)) \ sint (- a) = (- sint a)" "((- (2 ^ (size a - 1)) \ (sint a * sint b)) \ (sint a * sint b \ (2 ^ (size a - 1) - 1))) \ sint (a * b) = (sint a * sint b)" "((- (2 ^ (size a - 1)) \ (sint a sdiv sint b)) \ (sint a sdiv sint b \ (2 ^ (size a - 1) - 1))) \ sint (a sdiv b) = (sint a sdiv sint b)" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) \ sint (a smod b) = (sint a smod sint b)" by (subst (asm) signed_arith_ineq_checks_to_eq; simp)+ lemma signed_mult_eq_checks_double_size: assumes mult_le: "(2 ^ (len_of TYPE ('a) - 1) + 1) ^ 2 \ (2 :: int) ^ (len_of TYPE ('b) - 1)" and le: "2 ^ (LENGTH('a) - 1) \ (2 :: int) ^ (len_of TYPE ('b) - 1)" shows "(sint (a :: 'a :: len word) * sint b = sint (a * b)) = (scast a * scast b = (scast (a * b) :: 'b :: len word))" proof - have P: "sbintrunc (size a - 1) (sint a * sint b) \ range (sbintrunc (size a - 1))" by simp have abs: "!! x :: 'a word. abs (sint x) < 2 ^ (size a - 1) + 1" apply (cut_tac x=x in sint_range') apply (simp add: abs_le_iff word_size) done have abs_ab: "abs (sint a * sint b) < 2 ^ (LENGTH('b) - 1)" using abs_mult_less[OF abs[where x=a] abs[where x=b]] mult_le by (simp add: abs_mult power2_eq_square word_size) show ?thesis using P[unfolded range_sbintrunc] abs_ab le apply (simp add: sint_word_ariths scast_def) apply (simp add: wi_hom_mult) apply (subst word_sint.Abs_inject, simp_all) apply (simp add: sints_def range_sbintrunc abs_less_iff) apply clarsimp apply (simp add: sints_def range_sbintrunc word_size) apply (auto elim: order_less_le_trans order_trans[rotated]) done qed (* Properties about signed division. *) lemma int_sdiv_simps [simp]: "(a :: int) sdiv 1 = a" "(a :: int) sdiv 0 = 0" "(a :: int) sdiv -1 = -a" apply (auto simp: sdiv_int_def sgn_if) done lemma sgn_div_eq_sgn_mult: "a div b \ 0 \ sgn ((a :: int) div b) = sgn (a * b)" apply (clarsimp simp: sgn_if zero_le_mult_iff neg_imp_zdiv_nonneg_iff not_less) apply (metis less_le mult_le_0_iff neg_imp_zdiv_neg_iff not_less pos_imp_zdiv_neg_iff zdiv_eq_0_iff) done lemma sgn_sdiv_eq_sgn_mult: "a sdiv b \ 0 \ sgn ((a :: int) sdiv b) = sgn (a * b)" by (auto simp: sdiv_int_def sgn_div_eq_sgn_mult sgn_mult) lemma int_sdiv_same_is_1 [simp]: "a \ 0 \ ((a :: int) sdiv b = a) = (b = 1)" apply (rule iffI) apply (clarsimp simp: sdiv_int_def) apply (subgoal_tac "b > 0") apply (case_tac "a > 0") apply (clarsimp simp: sgn_if) apply (clarsimp simp: algebra_split_simps not_less) apply (metis int_div_same_is_1 le_neq_trans minus_minus neg_0_le_iff_le neg_equal_0_iff_equal) apply (case_tac "a > 0") apply (case_tac "b = 0") apply clarsimp apply (rule classical) apply (clarsimp simp: sgn_mult not_less) apply (metis le_less neg_0_less_iff_less not_less_iff_gr_or_eq pos_imp_zdiv_neg_iff) apply (rule classical) apply (clarsimp simp: algebra_split_simps sgn_mult not_less sgn_if split: if_splits) apply (metis antisym less_le neg_imp_zdiv_nonneg_iff) apply (clarsimp simp: sdiv_int_def sgn_if) done lemma int_sdiv_negated_is_minus1 [simp]: "a \ 0 \ ((a :: int) sdiv b = - a) = (b = -1)" apply (clarsimp simp: sdiv_int_def) apply (rule iffI) apply (subgoal_tac "b < 0") apply (case_tac "a > 0") apply (clarsimp simp: sgn_if algebra_split_simps not_less) apply (case_tac "sgn (a * b) = -1") apply (clarsimp simp: not_less algebra_split_simps) apply (clarsimp simp: algebra_split_simps not_less) apply (rule classical) apply (case_tac "b = 0") apply (clarsimp simp: not_less sgn_mult) apply (case_tac "a > 0") apply (clarsimp simp: not_less sgn_mult) apply (metis less_le neg_less_0_iff_less not_less_iff_gr_or_eq pos_imp_zdiv_neg_iff) apply (clarsimp simp: not_less sgn_mult) apply (metis antisym_conv div_minus_right neg_imp_zdiv_nonneg_iff neg_le_0_iff_le not_less) apply (clarsimp simp: sgn_if) done lemma sdiv_int_range: "(a :: int) sdiv b \ { - (abs a) .. (abs a) }" apply (unfold sdiv_int_def) apply (subgoal_tac "(abs a) div (abs b) \ (abs a)") apply (clarsimp simp: sgn_if) apply (meson abs_ge_zero neg_le_0_iff_le nonneg_mod_div order_trans) apply (metis abs_eq_0 abs_ge_zero div_by_0 zdiv_le_dividend zero_less_abs_iff) done lemma word_sdiv_div1 [simp]: "(a :: ('a::len) word) sdiv 1 = a" apply (rule sint_1_cases [where a=a]) apply (clarsimp simp: sdiv_word_def sdiv_int_def) apply (clarsimp simp: sdiv_word_def sdiv_int_def simp del: sint_minus1) apply (clarsimp simp: sdiv_word_def) done lemma sdiv_int_div_0 [simp]: "(x :: int) sdiv 0 = 0" by (clarsimp simp: sdiv_int_def) lemma sdiv_int_0_div [simp]: "0 sdiv (x :: int) = 0" by (clarsimp simp: sdiv_int_def) lemma word_sdiv_div0 [simp]: "(a :: ('a::len) word) sdiv 0 = 0" apply (auto simp: sdiv_word_def sdiv_int_def sgn_if) done lemma word_sdiv_div_minus1 [simp]: "(a :: ('a::len) word) sdiv -1 = -a" apply (auto simp: sdiv_word_def sdiv_int_def sgn_if) apply (metis wi_hom_neg word_sint.Rep_inverse') done lemmas word_sdiv_0 = word_sdiv_div0 lemma sdiv_word_min: "- (2 ^ (size a - 1)) \ sint (a :: ('a::len) word) sdiv sint (b :: ('a::len) word)" apply (clarsimp simp: word_size) apply (cut_tac sint_range' [where x=a]) apply (cut_tac sint_range' [where x=b]) apply clarsimp apply (insert sdiv_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: max_def abs_if split: if_split_asm) done lemma sdiv_word_max: "(sint (a :: ('a::len) word) sdiv sint (b :: ('a::len) word) < (2 ^ (size a - 1))) = ((a \ - (2 ^ (size a - 1)) \ (b \ -1)))" (is "?lhs = (\ ?a_int_min \ \ ?b_minus1)") proof (rule classical) assume not_thesis: "\ ?thesis" have not_zero: "b \ 0" using not_thesis by (clarsimp) have result_range: "sint a sdiv sint b \ (sints (size a)) \ {2 ^ (size a - 1)}" apply (cut_tac sdiv_int_range [where a="sint a" and b="sint b"]) apply (erule rev_subsetD) using sint_range' [where x=a] sint_range' [where x=b] apply (auto simp: max_def abs_if word_size sints_num) done have result_range_overflow: "(sint a sdiv sint b = 2 ^ (size a - 1)) = (?a_int_min \ ?b_minus1)" apply (rule iffI [rotated]) apply (clarsimp simp: sdiv_int_def sgn_if word_size sint_int_min) apply (rule classical) apply (case_tac "?a_int_min") apply (clarsimp simp: word_size sint_int_min) apply (metis diff_0_right int_sdiv_negated_is_minus1 minus_diff_eq minus_int_code(2) power_eq_0_iff sint_minus1 zero_neq_numeral) apply (subgoal_tac "abs (sint a) < 2 ^ (size a - 1)") apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] apply (clarsimp simp: word_size) apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] apply (insert word_sint.Rep [where x="a"])[1] apply (clarsimp simp: minus_le_iff word_size abs_if sints_num split: if_split_asm) apply (metis minus_minus sint_int_min word_sint.Rep_inject) done have result_range_simple: "(sint a sdiv sint b \ (sints (size a))) \ ?thesis" apply (insert sdiv_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: word_size sints_num sint_int_min) done show ?thesis apply (rule UnE [OF result_range result_range_simple]) apply simp apply (clarsimp simp: word_size) using result_range_overflow apply (clarsimp simp: word_size) done qed lemmas sdiv_word_min' = sdiv_word_min [simplified word_size, simplified] lemmas sdiv_word_max' = sdiv_word_max [simplified word_size, simplified] lemmas word_sdiv_numerals_lhs = sdiv_word_def[where a="numeral x" for x] sdiv_word_def[where a=0] sdiv_word_def[where a=1] lemmas word_sdiv_numerals = word_sdiv_numerals_lhs[where b="numeral y" for y] word_sdiv_numerals_lhs[where b=0] word_sdiv_numerals_lhs[where b=1] (* * Signed modulo properties. *) lemma smod_int_alt_def: "(a::int) smod b = sgn (a) * (abs a mod abs b)" apply (clarsimp simp: smod_int_def sdiv_int_def) apply (clarsimp simp: minus_div_mult_eq_mod [symmetric] abs_sgn sgn_mult sgn_if algebra_split_simps) done lemma smod_int_range: "b \ 0 \ (a::int) smod b \ { - abs b + 1 .. abs b - 1 }" apply (case_tac "b > 0") apply (insert pos_mod_conj [where a=a and b=b])[1] apply (insert pos_mod_conj [where a="-a" and b=b])[1] apply (auto simp: smod_int_alt_def algebra_simps sgn_if abs_if not_less add1_zle_eq [simplified add.commute])[1] apply (metis add_nonneg_nonneg int_one_le_iff_zero_less le_less less_add_same_cancel2 not_le pos_mod_conj) apply (metis (full_types) add.inverse_inverse eucl_rel_int eucl_rel_int_iff le_less_trans neg_0_le_iff_le) apply (insert neg_mod_conj [where a=a and b="b"])[1] apply (insert neg_mod_conj [where a="-a" and b="b"])[1] apply (clarsimp simp: smod_int_alt_def algebra_simps sgn_if abs_if not_less add1_zle_eq [simplified add.commute]) apply (metis neg_0_less_iff_less neg_mod_conj not_le not_less_iff_gr_or_eq order_trans pos_mod_conj) done lemma smod_int_compares: "\ 0 \ a; 0 < b \ \ (a :: int) smod b < b" "\ 0 \ a; 0 < b \ \ 0 \ (a :: int) smod b" "\ a \ 0; 0 < b \ \ -b < (a :: int) smod b" "\ a \ 0; 0 < b \ \ (a :: int) smod b \ 0" "\ 0 \ a; b < 0 \ \ (a :: int) smod b < - b" "\ 0 \ a; b < 0 \ \ 0 \ (a :: int) smod b" "\ a \ 0; b < 0 \ \ (a :: int) smod b \ 0" "\ a \ 0; b < 0 \ \ b \ (a :: int) smod b" apply (insert smod_int_range [where a=a and b=b]) apply (auto simp: add1_zle_eq smod_int_alt_def sgn_if) done lemma smod_int_mod_0 [simp]: "x smod (0 :: int) = x" by (clarsimp simp: smod_int_def) lemma smod_int_0_mod [simp]: "0 smod (x :: int) = 0" by (clarsimp simp: smod_int_alt_def) lemma smod_word_mod_0 [simp]: "x smod (0 :: ('a::len) word) = x" by (clarsimp simp: smod_word_def) lemma smod_word_0_mod [simp]: "0 smod (x :: ('a::len) word) = 0" by (clarsimp simp: smod_word_def) lemma smod_word_max: "sint (a::'a word) smod sint (b::'a word) < 2 ^ (LENGTH('a::len) - Suc 0)" apply (case_tac "b = 0") apply (insert word_sint.Rep [where x=a, simplified sints_num])[1] apply (clarsimp) apply (insert word_sint.Rep [where x="b", simplified sints_num])[1] apply (insert smod_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: abs_if split: if_split_asm) done lemma smod_word_min: "- (2 ^ (LENGTH('a::len) - Suc 0)) \ sint (a::'a word) smod sint (b::'a word)" apply (case_tac "b = 0") apply (insert word_sint.Rep [where x=a, simplified sints_num])[1] apply clarsimp apply (insert word_sint.Rep [where x=b, simplified sints_num])[1] apply (insert smod_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: abs_if add1_zle_eq split: if_split_asm) done lemma smod_word_alt_def: "(a :: ('a::len) word) smod b = a - (a sdiv b) * b" apply (case_tac "a \ - (2 ^ (LENGTH('a) - 1)) \ b \ -1") apply (clarsimp simp: smod_word_def sdiv_word_def smod_int_def minus_word.abs_eq [symmetric] times_word.abs_eq [symmetric]) apply (clarsimp simp: smod_word_def smod_int_def) done lemmas word_smod_numerals_lhs = smod_word_def[where a="numeral x" for x] smod_word_def[where a=0] smod_word_def[where a=1] lemmas word_smod_numerals = word_smod_numerals_lhs[where b="numeral y" for y] word_smod_numerals_lhs[where b=0] word_smod_numerals_lhs[where b=1] lemma sint_of_int_eq: "\ - (2 ^ (LENGTH('a) - 1)) \ x; x < 2 ^ (LENGTH('a) - 1) \ \ sint (of_int x :: ('a::len) word) = x" apply (clarsimp simp: word_of_int int_word_sint) apply (subst int_mod_eq') apply simp apply (subst (2) power_minus_simp) apply clarsimp apply clarsimp apply clarsimp done lemma of_int_sint [simp]: "of_int (sint a) = a" apply (insert word_sint.Rep [where x=a]) apply (clarsimp simp: word_of_int) done lemma nth_w2p_scast [simp]: "((scast ((2::'a::len signed word) ^ n) :: 'a word) !! m) \ ((((2::'a::len word) ^ n) :: 'a word) !! m)" apply (subst nth_w2p) apply (case_tac "n \ LENGTH('a)") apply (subst power_overflow, simp) apply clarsimp apply (metis nth_w2p scast_def test_bit_conj_lt len_signed nth_word_of_int word_sint.Rep_inverse) done lemma scast_2_power [simp]: "scast ((2 :: 'a::len signed word) ^ x) = ((2 :: 'a word) ^ x)" by (clarsimp simp: word_eq_iff) lemma scast_bit_test [simp]: "scast ((1 :: 'a::len signed word) << n) = (1 :: 'a word) << n" by (clarsimp simp: word_eq_iff) lemma ucast_nat_def': "of_nat (unat x) = (ucast :: 'a :: len word \ ('b :: len) signed word) x" by (simp add: ucast_def word_of_int_nat unat_def) lemma mod_mod_power_int: fixes k :: int shows "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ (min m n)" by (metis bintrunc_bintrunc_min bintrunc_mod2p min.commute) (* Normalise combinations of scast and ucast. *) lemma ucast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" fixes L :: "int \ int \ int" assumes lift_M: "\x y. uint (M x y) = L (uint x) (uint y) mod 2 ^ LENGTH('a)" assumes lift_M': "\x y. uint (M' x y) = L (uint x) (uint y) mod 2 ^ LENGTH('b)" assumes distrib: "\x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b)) = (L x y) mod (2 ^ LENGTH('b))" assumes is_down: "is_down (ucast :: 'a word \ 'b word)" shows "ucast (M a b) = M' (ucast a) (ucast b)" apply (clarsimp simp: word_of_int ucast_def) apply (subst lift_M) apply (subst of_int_uint [symmetric], subst lift_M') apply (subst (1 2) int_word_uint) apply (subst word_of_int) apply (subst word.abs_eq_iff) apply (subst (1 2) bintrunc_mod2p) apply (insert is_down) apply (unfold is_down_def) apply (clarsimp simp: target_size source_size) apply (clarsimp simp: mod_mod_power_int min_def) apply (rule distrib [symmetric]) done lemma ucast_down_add: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) + b) = (ucast a + ucast b :: 'b::len word)" by (rule ucast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp) lemma ucast_down_minus: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) - b) = (ucast a - ucast b :: 'b::len word)" apply (rule ucast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_diff_left_eq mod_diff_right_eq) apply simp done lemma ucast_down_mult: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) * b) = (ucast a * ucast b :: 'b::len word)" apply (rule ucast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_mult_eq) apply simp done lemma scast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" fixes L :: "int \ int \ int" assumes lift_M: "\x y. uint (M x y) = L (uint x) (uint y) mod 2 ^ LENGTH('a)" assumes lift_M': "\x y. uint (M' x y) = L (uint x) (uint y) mod 2 ^ LENGTH('b)" assumes distrib: "\x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b)) = (L x y) mod (2 ^ LENGTH('b))" assumes is_down: "is_down (scast :: 'a word \ 'b word)" shows "scast (M a b) = M' (scast a) (scast b)" apply (subst (1 2 3) down_cast_same [symmetric]) apply (insert is_down) apply (clarsimp simp: is_down_def target_size source_size is_down) apply (rule ucast_distrib [where L=L, OF lift_M lift_M' distrib]) apply (insert is_down) apply (clarsimp simp: is_down_def target_size source_size is_down) done lemma scast_down_add: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) + b) = (scast a + scast b :: 'b::len word)" by (rule scast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp) lemma scast_down_minus: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) - b) = (scast a - scast b :: 'b::len word)" apply (rule scast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_diff_left_eq mod_diff_right_eq) apply simp done lemma scast_down_mult: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) * b) = (scast a * scast b :: 'b::len word)" apply (rule scast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_mult_eq) apply simp done lemma scast_ucast_1: "\ is_down (ucast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_def ucast_down_wi) lemma scast_ucast_3: "\ is_down (ucast :: 'a word \ 'c word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_def ucast_down_wi) lemma scast_ucast_4: "\ is_up (ucast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_def ucast_down_wi) lemma scast_scast_b: "\ is_up (scast :: 'a word \ 'b word) \ \ (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_def sint_up_scast) lemma ucast_scast_1: "\ is_down (scast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_def ucast_down_wi) lemma ucast_scast_3: "\ is_down (scast :: 'a word \ 'c word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_def ucast_down_wi) lemma ucast_scast_4: "\ is_up (scast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis down_cast_same scast_def sint_up_scast) lemma ucast_ucast_a: "\ is_down (ucast :: 'b word \ 'c word) \ \ (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_def ucast_down_wi) lemma ucast_ucast_b: "\ is_up (ucast :: 'a word \ 'b word) \ \ (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis ucast_up_ucast) lemma scast_scast_a: "\ is_down (scast :: 'b word \ 'c word) \ \ (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" apply (clarsimp simp: scast_def) apply (metis down_cast_same is_up_down scast_def ucast_down_wi) done lemma scast_down_wi [OF refl]: "uc = scast \ is_down uc \ uc (word_of_int x) = word_of_int x" by (metis down_cast_same is_up_down ucast_down_wi) lemmas cast_simps = is_down is_up scast_down_add scast_down_minus scast_down_mult ucast_down_add ucast_down_minus ucast_down_mult scast_ucast_1 scast_ucast_3 scast_ucast_4 ucast_scast_1 ucast_scast_3 ucast_scast_4 ucast_ucast_a ucast_ucast_b scast_scast_a scast_scast_b ucast_down_bl ucast_down_wi scast_down_wi ucast_of_nat scast_of_nat uint_up_ucast sint_up_scast up_scast_surj up_ucast_surj lemma smod_mod_positive: "\ 0 \ (a :: int); 0 \ b \ \ a smod b = a mod b" by (clarsimp simp: smod_int_alt_def zsgn_def) lemma nat_mult_power_less_eq: "b > 0 \ (a * b ^ n < (b :: nat) ^ m) = (a < b ^ (m - n))" using mult_less_cancel2[where m = a and k = "b ^ n" and n="b ^ (m - n)"] mult_less_cancel2[where m="a * b ^ (n - m)" and k="b ^ m" and n=1] apply (simp only: power_add[symmetric] nat_minus_add_max) apply (simp only: power_add[symmetric] nat_minus_add_max ac_simps) apply (simp add: max_def split: if_split_asm) done lemma signed_shift_guard_to_word: "\ n < len_of TYPE ('a); n > 0 \ \ (unat (x :: 'a :: len word) * 2 ^ y < 2 ^ n) = (x = 0 \ x < (1 << n >> y))" apply (simp only: nat_mult_power_less_eq) apply (cases "y \ n") apply (simp only: shiftl_shiftr1) apply (subst less_mask_eq) apply (simp add: word_less_nat_alt word_size) apply (rule order_less_le_trans[rotated], rule power_increasing[where n=1]) apply simp apply simp apply simp apply (simp add: nat_mult_power_less_eq word_less_nat_alt word_size) apply auto[1] apply (simp only: shiftl_shiftr2, simp add: unat_eq_0) done lemma sint_ucast_eq_uint: "\ \ is_down (ucast :: ('a::len word \ 'b::len word)) \ \ sint ((ucast :: ('a::len word \ 'b::len word)) x) = uint x" apply (subst sint_eq_uint) apply (clarsimp simp: msb_nth nth_ucast is_down) apply (metis Suc_leI Suc_pred len_gt_0) apply (clarsimp simp: uint_up_ucast is_up is_down) done lemma word_less_nowrapI': - "(x :: 'a :: len0 word) \ z - k \ k \ z \ 0 < k \ x < x + k" + "(x :: 'a :: len word) \ z - k \ k \ z \ 0 < k \ x < x + k" by uint_arith lemma mask_plus_1: "mask n + 1 = 2 ^ n" by (clarsimp simp: mask_def) lemma unat_inj: "inj unat" by (metis eq_iff injI word_le_nat_alt) lemma unat_ucast_upcast: "is_up (ucast :: 'b word \ 'a word) \ unat (ucast x :: ('a::len) word) = unat (x :: ('b::len) word)" unfolding ucast_def unat_def apply (subst int_word_uint) apply (subst mod_pos_pos_trivial) apply simp apply (rule lt2p_lem) apply (clarsimp simp: is_up) apply simp done lemma ucast_mono: "\ (x :: 'b :: len word) < y; y < 2 ^ LENGTH('a) \ \ ucast x < ((ucast y) :: 'a :: len word)" apply (simp add: ucast_nat_def [symmetric]) apply (rule of_nat_mono_maybe) apply (rule unat_less_helper) apply (simp add: Power.of_nat_power) apply (simp add: word_less_nat_alt) done lemma ucast_mono_le: "\x \ y; y < 2 ^ LENGTH('b)\ \ (ucast (x :: 'a :: len word) :: 'b :: len word) \ ucast y" apply (simp add: ucast_nat_def [symmetric]) apply (subst of_nat_mono_maybe_le[symmetric]) apply (rule unat_less_helper) apply (simp add: Power.of_nat_power) apply (rule unat_less_helper) apply (erule le_less_trans) apply (simp add: Power.of_nat_power) apply (simp add: word_le_nat_alt) done lemma ucast_mono_le': "\ unat y < 2 ^ LENGTH('b); LENGTH('b::len) < LENGTH('a::len); x \ y \ \ UCAST('a \ 'b) x \ UCAST('a \ 'b) y" by (auto simp: word_less_nat_alt intro: ucast_mono_le) lemma zero_sle_ucast_up: "\ is_down (ucast :: 'a word \ 'b signed word) \ (0 <=s ((ucast (b::('a::len) word)) :: ('b::len) signed word))" apply (subgoal_tac "\ msb (ucast b :: 'b signed word)") apply (clarsimp simp: word_sle_msb_le) apply (clarsimp simp: is_down not_le msb_nth nth_ucast) apply (subst (asm) test_bit_conj_lt [symmetric]) apply clarsimp apply arith done lemma word_le_ucast_sless: "\ x \ y; y \ -1; LENGTH('a) < LENGTH('b) \ \ UCAST (('a :: len) \ ('b :: len) signed) x msb (ucast x :: ('a::len) word) = msb (x :: ('b::len) word)" apply (clarsimp simp: word_msb_alt) apply (subst ucast_down_drop [where n=0]) apply (clarsimp simp: source_size_def target_size_def word_size) apply clarsimp done lemma msb_big: "msb (a :: ('a::len) word) = (a \ 2 ^ (LENGTH('a) - Suc 0))" apply (rule iffI) apply (clarsimp simp: msb_nth) apply (drule bang_is_le) apply simp apply (rule ccontr) apply (subgoal_tac "a = a && mask (LENGTH('a) - Suc 0)") apply (cut_tac and_mask_less' [where w=a and n="LENGTH('a) - Suc 0"]) apply (clarsimp simp: word_not_le [symmetric]) apply clarsimp apply (rule sym, subst and_mask_eq_iff_shiftr_0) apply (clarsimp simp: msb_shift) done lemma zero_sle_ucast: "(0 <=s ((ucast (b::('a::len) word)) :: ('a::len) signed word)) = (uint b < 2 ^ (LENGTH('a) - 1))" apply (case_tac "msb b") apply (clarsimp simp: word_sle_msb_le not_less msb_ucast_eq del: notI) apply (clarsimp simp: msb_big word_le_def uint_2p_alt) apply (clarsimp simp: word_sle_msb_le not_less msb_ucast_eq del: notI) apply (clarsimp simp: msb_big word_le_def uint_2p_alt) done (* to_bool / from_bool. *) definition from_bool :: "bool \ 'a::len word" where "from_bool b \ case b of True \ of_nat 1 | False \ of_nat 0" lemma from_bool_0: "(from_bool x = 0) = (\ x)" by (simp add: from_bool_def split: bool.split) definition to_bool :: "'a::len word \ bool" where "to_bool \ (\) 0" lemma to_bool_and_1: "to_bool (x && 1) = (x !! 0)" apply (simp add: to_bool_def) apply (rule iffI) apply (rule classical, erule notE, word_eqI) apply (case_tac n; simp) apply (rule notI, drule word_eqD[where x=0]) apply simp done lemma to_bool_from_bool [simp]: "to_bool (from_bool r) = r" unfolding from_bool_def to_bool_def by (simp split: bool.splits) lemma from_bool_neq_0 [simp]: "(from_bool b \ 0) = b" by (simp add: from_bool_def split: bool.splits) lemma from_bool_mask_simp [simp]: "(from_bool r :: 'a::len word) && 1 = from_bool r" unfolding from_bool_def by (clarsimp split: bool.splits) lemma from_bool_1 [simp]: "(from_bool P = 1) = P" by (simp add: from_bool_def split: bool.splits) lemma ge_0_from_bool [simp]: "(0 < from_bool P) = P" by (simp add: from_bool_def split: bool.splits) lemma limited_and_from_bool: "limited_and (from_bool b) 1" by (simp add: from_bool_def limited_and_def split: bool.split) lemma to_bool_1 [simp]: "to_bool 1" by (simp add: to_bool_def) lemma to_bool_0 [simp]: "\to_bool 0" by (simp add: to_bool_def) lemma from_bool_eq_if: "(from_bool Q = (if P then 1 else 0)) = (P = Q)" by (simp add: case_bool_If from_bool_def split: if_split) lemma to_bool_eq_0: "(\ to_bool x) = (x = 0)" by (simp add: to_bool_def) lemma to_bool_neq_0: "(to_bool x) = (x \ 0)" by (simp add: to_bool_def) lemma from_bool_all_helper: "(\bool. from_bool bool = val \ P bool) = ((\bool. from_bool bool = val) \ P (val \ 0))" by (auto simp: from_bool_0) lemma from_bool_to_bool_iff: "w = from_bool b \ to_bool w = b \ (w = 0 \ w = 1)" by (cases b) (auto simp: from_bool_def to_bool_def) lemma from_bool_eqI: "from_bool x = from_bool y \ x = y" unfolding from_bool_def by (auto split: bool.splits) lemma word_rsplit_upt: "\ size x = LENGTH('a :: len) * n; n \ 0 \ \ word_rsplit x = map (\i. ucast (x >> i * len_of TYPE ('a)) :: 'a word) (rev [0 ..< n])" apply (subgoal_tac "length (word_rsplit x :: 'a word list) = n") apply (rule nth_equalityI, simp) apply (intro allI word_eqI impI) apply (simp add: test_bit_rsplit_alt word_size) apply (simp add: nth_ucast nth_shiftr nth_rev field_simps) apply (simp add: length_word_rsplit_exp_size) apply (metis mult.commute given_quot_alt word_size word_size_gt_0) done lemma aligned_shift: "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ \ x + y >> n = y >> n" by (subst word_plus_and_or_coroll; word_eqI, blast) lemma aligned_shift': "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ \ y + x >> n = y >> n" by (subst word_plus_and_or_coroll; word_eqI, blast) lemma neg_mask_add_mask: "((x:: 'a :: len word) && ~~ (mask n)) + (2 ^ n - 1) = x || mask n" unfolding mask_2pm1[symmetric] by (subst word_plus_and_or_coroll; word_eqI_solve) lemma subtract_mask: "p - (p && mask n) = (p && ~~ (mask n))" "p - (p && ~~ (mask n)) = (p && mask n)" by (simp add: field_simps word_plus_and_or_coroll2)+ lemma and_neg_mask_plus_mask_mono: "(p && ~~ (mask n)) + mask n \ p" apply (rule word_le_minus_cancel[where x = "p && ~~ (mask n)"]) apply (clarsimp simp: subtract_mask) using word_and_le1[where a = "mask n" and y = p] apply (clarsimp simp: mask_def word_le_less_eq) apply (rule is_aligned_no_overflow'[folded mask_2pm1]) apply (clarsimp simp: is_aligned_neg_mask) done lemma word_neg_and_le: "ptr \ (ptr && ~~ (mask n)) + (2 ^ n - 1)" by (simp add: and_neg_mask_plus_mask_mono mask_2pm1[symmetric]) lemma aligned_less_plus_1: "\ is_aligned x n; n > 0 \ \ x < x + 1" apply (rule plus_one_helper2) apply (rule order_refl) apply (clarsimp simp: field_simps) apply (drule arg_cong[where f="\x. x - 1"]) apply (clarsimp simp: is_aligned_mask) apply (drule word_eqD[where x=0]) apply simp done lemma aligned_add_offset_less: "\is_aligned x n; is_aligned y n; x < y; z < 2 ^ n\ \ x + z < y" apply (cases "y = 0") apply simp apply (erule is_aligned_get_word_bits[where p=y], simp_all) apply (cases "z = 0", simp_all) apply (drule(2) aligned_at_least_t2n_diff[rotated -1]) apply (drule plus_one_helper2) apply (rule less_is_non_zero_p1) apply (rule aligned_less_plus_1) apply (erule aligned_sub_aligned[OF _ _ order_refl], simp_all add: is_aligned_triv)[1] apply (cases n, simp_all)[1] apply (simp only: trans[OF diff_add_eq diff_diff_eq2[symmetric]]) apply (drule word_less_add_right) apply (rule ccontr, simp add: linorder_not_le) apply (drule aligned_small_is_0, erule order_less_trans) apply (clarsimp simp: power_overflow) apply simp apply (erule order_le_less_trans[rotated], rule word_plus_mono_right) apply (erule word_le_minus_one_leq) apply (simp add: is_aligned_no_wrap' is_aligned_no_overflow field_simps) done lemma is_aligned_add_helper: "\ is_aligned p n; d < 2 ^ n \ \ (p + d && mask n = d) \ (p + d && (~~ (mask n)) = p)" apply (subst(asm) is_aligned_mask) apply (drule less_mask_eq) apply (rule context_conjI) apply (subst word_plus_and_or_coroll; word_eqI; blast) using word_plus_and_or_coroll2[where x="p + d" and w="mask n"] by simp lemma is_aligned_sub_helper: "\ is_aligned (p - d) n; d < 2 ^ n \ \ (p && mask n = d) \ (p && (~~ (mask n)) = p - d)" by (drule(1) is_aligned_add_helper, simp) lemma mask_twice: "(x && mask n) && mask m = x && mask (min m n)" by word_eqI_solve lemma is_aligned_after_mask: "\is_aligned k m;m\ n\ \ is_aligned (k && mask n) m" by (rule is_aligned_andI1) lemma and_mask_plus: "\is_aligned ptr m; m \ n; a < 2 ^ m\ \ ptr + a && mask n = (ptr && mask n) + a" apply (rule mask_eqI[where n = m]) apply (simp add:mask_twice min_def) apply (simp add:is_aligned_add_helper) apply (subst is_aligned_add_helper[THEN conjunct1]) apply (erule is_aligned_after_mask) apply simp apply simp apply simp apply (subgoal_tac "(ptr + a && mask n) && ~~ (mask m) = (ptr + a && ~~ (mask m) ) && mask n") apply (simp add:is_aligned_add_helper) apply (subst is_aligned_add_helper[THEN conjunct2]) apply (simp add:is_aligned_after_mask) apply simp apply simp apply (simp add:word_bw_comms word_bw_lcs) done lemma le_step_down_word:"\(i::('a::len) word) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by unat_arith lemma le_step_down_word_2: fixes x :: "'a::len word" shows "\x \ y; x \ y\ \ x \ y - 1" by (subst (asm) word_le_less_eq, clarsimp, simp add: word_le_minus_one_leq) lemma NOT_mask_AND_mask[simp]: "(w && mask n) && ~~ (mask n) = 0" apply (clarsimp simp:mask_def) by (metis word_bool_alg.conj_cancel_right word_bool_alg.conj_zero_right word_bw_comms(1) word_bw_lcs(1)) lemma and_and_not[simp]:"(a && b) && ~~ b = 0" apply (subst word_bw_assocs(1)) apply clarsimp done lemma mask_shift_and_negate[simp]:"(w && mask n << m) && ~~ (mask n << m) = 0" apply (clarsimp simp:mask_def) by (metis (erased, hide_lams) mask_eq_x_eq_0 shiftl_over_and_dist word_bool_alg.conj_absorb word_bw_assocs(1)) lemma le_step_down_nat:"\(i::nat) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by arith lemma le_step_down_int:"\(i::int) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by arith lemma ex_mask_1[simp]: "(\x. mask x = 1)" apply (rule_tac x=1 in exI) apply (simp add:mask_def) done lemma not_switch:"~~ a = x \ a = ~~ x" by auto (* The seL4 bitfield generator produces functions containing mask and shift operations, such that * invoking two of them consecutively can produce something like the following. *) lemma bitfield_op_twice: "(x && ~~ (mask n << m) || ((y && mask n) << m)) && ~~ (mask n << m) = x && ~~ (mask n << m)" by (induct n arbitrary: m) (auto simp: word_ao_dist) lemma bitfield_op_twice'': "\~~ a = b << c; \x. b = mask x\ \ (x && a || (y && b << c)) && a = x && a" apply clarsimp apply (cut_tac n=xa and m=c and x=x and y=y in bitfield_op_twice) apply (clarsimp simp:mask_def) apply (drule not_switch) apply clarsimp done lemma bit_twiddle_min: "(y::'a::len word) xor (((x::'a::len word) xor y) && (if x < y then -1 else 0)) = min x y" by (metis (mono_tags) min_def word_bitwise_m1_simps(2) word_bool_alg.xor_left_commute word_bool_alg.xor_self word_bool_alg.xor_zero_right word_bw_comms(1) word_le_less_eq word_log_esimps(7)) lemma bit_twiddle_max: "(x::'a::len word) xor (((x::'a::len word) xor y) && (if x < y then -1 else 0)) = max x y" by (metis (mono_tags) max_def word_bitwise_m1_simps(2) word_bool_alg.xor_left_self word_bool_alg.xor_zero_right word_bw_comms(1) word_le_less_eq word_log_esimps(7)) lemma swap_with_xor: "\(x::'a::len word) = a xor b; y = b xor x; z = x xor y\ \ z = b \ y = a" by (metis word_bool_alg.xor_assoc word_bool_alg.xor_commute word_bool_alg.xor_self word_bool_alg.xor_zero_right) lemma scast_nop1: "((scast ((of_int x)::('a::len) word))::'a sword) = of_int x" apply (clarsimp simp:scast_def word_of_int) by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse) lemma scast_nop2: "((scast ((of_int x)::('a::len) sword))::'a word) = of_int x" apply (clarsimp simp:scast_def word_of_int) by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse) lemmas scast_nop[simp] = scast_nop1 scast_nop2 scast_id lemma le_mask_imp_and_mask: "(x::'a::len word) \ mask n \ x && mask n = x" by (metis and_mask_eq_iff_le_mask) lemma or_not_mask_nop: "((x::'a::len word) || ~~ (mask n)) && mask n = x && mask n" by (metis word_and_not word_ao_dist2 word_bw_comms(1) word_log_esimps(3)) lemma mask_subsume: "\n \ m\ \ ((x::'a::len word) || y && mask n) && ~~ (mask m) = x && ~~ (mask m)" apply (subst word_ao_dist) apply (subgoal_tac "(y && mask n) && ~~ (mask m) = 0") apply simp by (metis (no_types, hide_lams) is_aligned_mask is_aligned_weaken word_and_not word_bool_alg.conj_zero_right word_bw_comms(1) word_bw_lcs(1)) lemma and_mask_0_iff_le_mask: fixes w :: "'a::len word" shows "(w && ~~(mask n) = 0) = (w \ mask n)" by (simp add: mask_eq_0_eq_x le_mask_imp_and_mask and_mask_eq_iff_le_mask) lemma mask_twice2: "n \ m \ ((x::'a::len word) && mask m) && mask n = x && mask n" by (metis mask_twice min_def) lemma uint_2_id: "LENGTH('a) \ 2 \ uint (2::('a::len) word) = 2" by simp lemma bintrunc_id: "\m \ of_nat n; 0 < m\ \ bintrunc n m = m" by (simp add: bintrunc_mod2p le_less_trans) lemma shiftr1_unfold: "shiftr1 x = x >> 1" by (metis One_nat_def comp_apply funpow.simps(1) funpow.simps(2) id_apply shiftr_def) lemma shiftr1_is_div_2: "(x::('a::len) word) >> 1 = x div 2" proof (cases \2 \ LENGTH('a)\) case False then have *: \LENGTH('a) = 1\ by simp then have \x = 0 \ x = 1\ by (metis One_nat_def less_irrefl_nat sint_1_cases) then show ?thesis by (auto simp add: word_arith_nat_defs(6) *) next case True then show ?thesis using shiftr1_unfold [symmetric, of x] uint_2_id [where ?'a = 'a] by (simp add: shiftr1_def word_div_def) qed lemma shiftl1_is_mult: "(x << 1) = (x :: 'a::len word) * 2" by (metis One_nat_def mult_2 mult_2_right one_add_one power_0 power_Suc shiftl_t2n) lemma div_of_0_id[simp]:"(0::('a::len) word) div n = 0" by (simp add: word_div_def) lemma degenerate_word:"LENGTH('a) = 1 \ (x::('a::len) word) = 0 \ x = 1" by (metis One_nat_def less_irrefl_nat sint_1_cases) lemma div_by_0_word:"(x::('a::len) word) div 0 = 0" by (metis div_0 div_by_0 unat_0 word_arith_nat_defs(6) word_div_1) lemma div_less_dividend_word:"\x \ 0; n \ 1\ \ (x::('a::len) word) div n < x" - apply (case_tac "n = 0") + apply (cases \n = 0\) apply clarsimp - apply (subst div_by_0_word) apply (simp add:word_neq_0_conv) apply (subst word_arith_nat_div) apply (rule word_of_nat_less) apply (rule div_less_dividend) - apply (metis (poly_guards_query) One_nat_def less_one nat_neq_iff unat_eq_1(2) unat_eq_zero) + using unat_eq_zero word_unat_Rep_inject1 apply force apply (simp add:unat_gt_0) done lemma shiftr1_lt:"x \ 0 \ (x::('a::len) word) >> 1 < x" apply (subst shiftr1_is_div_2) apply (rule div_less_dividend_word) apply simp+ done lemma word_less_div: fixes x :: "('a::len) word" and y :: "('a::len) word" shows "x div y = 0 \ y = 0 \ x < y" apply (case_tac "y = 0", clarsimp+) by (metis One_nat_def Suc_le_mono le0 le_div_geq not_less unat_0 unat_div unat_gt_0 word_less_nat_alt zero_less_one) lemma not_degenerate_imp_2_neq_0:"LENGTH('a) > 1 \ (2::('a::len) word) \ 0" by (metis numerals(1) power_not_zero power_zero_numeral) lemma shiftr1_0_or_1:"(x::('a::len) word) >> 1 = 0 \ x = 0 \ x = 1" apply (subst (asm) shiftr1_is_div_2) apply (drule word_less_div) apply (case_tac "LENGTH('a) = 1") apply (simp add:degenerate_word) apply (erule disjE) apply (subgoal_tac "(2::'a word) \ 0") apply simp apply (rule not_degenerate_imp_2_neq_0) apply (subgoal_tac "LENGTH('a) \ 0") apply arith apply simp apply (rule x_less_2_0_1', simp+) done lemma word_overflow:"(x::('a::len) word) + 1 > x \ x + 1 = 0" apply clarsimp by (metis diff_0 eq_diff_eq less_x_plus_1) lemma word_overflow_unat:"unat ((x::('a::len) word) + 1) = unat x + 1 \ x + 1 = 0" by (metis Suc_eq_plus1 add.commute unatSuc) lemma even_word_imp_odd_next:"even (unat (x::('a::len) word)) \ x + 1 = 0 \ odd (unat (x + 1))" apply (cut_tac x=x in word_overflow_unat) apply clarsimp done lemma odd_word_imp_even_next:"odd (unat (x::('a::len) word)) \ x + 1 = 0 \ even (unat (x + 1))" apply (cut_tac x=x in word_overflow_unat) apply clarsimp done lemma overflow_imp_lsb:"(x::('a::len) word) + 1 = 0 \ x !! 0" by (metis add.commute add_left_cancel max_word_max not_less word_and_1 word_bool_alg.conj_one_right word_bw_comms(1) word_overflow zero_neq_one) lemma word_lsb_nat:"lsb w = (unat w mod 2 = 1)" unfolding word_lsb_def bin_last_def by (metis (no_types, hide_lams) nat_mod_distrib nat_numeral not_mod_2_eq_1_eq_0 numeral_One uint_eq_0 uint_nonnegative unat_0 unat_def zero_le_numeral) lemma odd_iff_lsb:"odd (unat (x::('a::len) word)) = x !! 0" apply (simp add:even_iff_mod_2_eq_zero) apply (subst word_lsb_nat[unfolded One_nat_def, symmetric]) apply (rule word_lsb_alt) done lemma of_nat_neq_iff_word: "x mod 2 ^ LENGTH('a) \ y mod 2 ^ LENGTH('a) \ (((of_nat x)::('a::len) word) \ of_nat y) = (x \ y)" apply (rule iffI) apply (case_tac "x = y") apply (subst (asm) of_nat_eq_iff[symmetric]) apply simp+ apply (case_tac "((of_nat x)::('a::len) word) = of_nat y") apply (subst (asm) word_unat.norm_eq_iff[symmetric]) apply simp+ done lemma shiftr1_irrelevant_lsb:"(x::('a::len) word) !! 0 \ x >> 1 = (x + 1) >> 1" using word_overflow_unat [of x] apply (simp only: shiftr1_is_div_2 flip: odd_iff_lsb) apply (cases \2 \ LENGTH('a)\) apply (auto simp add: test_bit_def' uint_nat word_arith_nat_div dest: overflow_imp_lsb) done lemma shiftr1_0_imp_only_lsb:"((x::('a::len) word) + 1) >> 1 = 0 \ x = 0 \ x + 1 = 0" by (metis One_nat_def shiftr1_0_or_1 word_less_1 word_overflow) lemma shiftr1_irrelevant_lsb':"\((x::('a::len) word) !! 0) \ x >> 1 = (x + 1) >> 1" by (metis shiftr1_irrelevant_lsb) lemma lsb_this_or_next:"\(((x::('a::len) word) + 1) !! 0) \ x !! 0" by (metis (poly_guards_query) even_word_imp_odd_next odd_iff_lsb overflow_imp_lsb) (* Perhaps this one should be a simp lemma, but it seems a little dangerous. *) lemma cast_chunk_assemble_id: "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ (((ucast ((ucast (x::'b word))::'a word))::'b word) || (((ucast ((ucast (x >> n))::'a word))::'b word) << n)) = x" apply (subgoal_tac "((ucast ((ucast (x >> n))::'a word))::'b word) = x >> n") apply clarsimp apply (subst and_not_mask[symmetric]) apply (subst ucast_ucast_mask) apply (subst word_ao_dist2[symmetric]) apply clarsimp apply (rule ucast_ucast_len) apply (rule shiftr_less_t2n') apply (subst and_mask_eq_iff_le_mask) apply (simp_all add: mask_def flip: mult_2_right) apply (metis add_diff_cancel_left' len_gt_0 mult_2_right zero_less_diff) done lemma cast_chunk_scast_assemble_id: "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ (((ucast ((scast (x::'b word))::'a word))::'b word) || (((ucast ((scast (x >> n))::'a word))::'b word) << n)) = x" apply (subgoal_tac "((scast x)::'a word) = ((ucast x)::'a word)") apply (subgoal_tac "((scast (x >> n))::'a word) = ((ucast (x >> n))::'a word)") apply (simp add:cast_chunk_assemble_id) apply (subst down_cast_same[symmetric], subst is_down, arith, simp)+ done lemma mask_or_not_mask: "x && mask n || x && ~~ (mask n) = x" apply (subst word_oa_dist, simp) apply (subst word_oa_dist2, simp) done lemma is_aligned_add_not_aligned: "\is_aligned (p::'a::len word) n; \ is_aligned (q::'a::len word) n\ \ \ is_aligned (p + q) n" by (metis is_aligned_addD1) lemma word_gr0_conv_Suc: "(m::'a::len word) > 0 \ \n. m = n + 1" by (metis add.commute add_minus_cancel) lemma neg_mask_add_aligned: "\ is_aligned p n; q < 2 ^ n \ \ (p + q) && ~~ (mask n) = p && ~~ (mask n)" by (metis is_aligned_add_helper is_aligned_neg_mask_eq) lemma word_sless_sint_le:"x sint x \ sint y - 1" by (metis word_sless_alt zle_diff1_eq) lemma upper_trivial: fixes x :: "'a::len word" shows "x \ 2 ^ LENGTH('a) - 1 \ x < 2 ^ LENGTH('a) - 1" by (simp add: less_le) lemma constraint_expand: fixes x :: "'a::len word" shows "x \ {y. lower \ y \ y \ upper} = (lower \ x \ x \ upper)" by (rule mem_Collect_eq) lemma card_map_elide: "card ((of_nat :: nat \ 'a::len word) ` {0.. CARD('a::len word)" proof - let ?of_nat = "of_nat :: nat \ 'a word" from word_unat.Abs_inj_on have "inj_on ?of_nat {i. i < CARD('a word)}" by (simp add: unats_def card_word) moreover have "{0.. {i. i < CARD('a word)}" using that by auto ultimately have "inj_on ?of_nat {0.. CARD('a::len word) \ card ((of_nat::nat \ 'a::len word) ` {0..UCAST('b \ 'a) (UCAST('a \ 'b) x) = x\ if \x \ UCAST('b::len \ 'a) (- 1)\ for x :: \'a::len word\ proof - from that have a1: \x \ word_of_int (uint (word_of_int (2 ^ LENGTH('b) - 1) :: 'b word))\ by (simp add: ucast_def word_of_int_minus) have f2: "((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ \ (0::int) \ - 1 + 2 ^ LENGTH('b) \ (0::int) \ - 1 + 2 ^ LENGTH('b) + - 1 * 2 ^ LENGTH('b) \ (- (1::int) + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b) = - 1 + 2 ^ LENGTH('b)) = ((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ \ (1::int) \ 2 ^ LENGTH('b) \ 2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)) = 1)" by force have f3: "\i ia. \ (0::int) \ i \ 0 \ i + - 1 * ia \ i mod ia = i" using mod_pos_pos_trivial by force have "(1::int) \ 2 ^ LENGTH('b)" by simp then have "2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ len_of TYPE ('b)) = 1" using f3 f2 by blast then have f4: "- (1::int) + 2 ^ LENGTH('b) = (- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)" by linarith have f5: "x \ word_of_int (uint (word_of_int (- 1 + 2 ^ LENGTH('b))::'b word))" using a1 by force have f6: "2 ^ LENGTH('b) + - (1::int) = - 1 + 2 ^ LENGTH('b)" by force have f7: "- (1::int) * 1 = - 1" by auto have "\x0 x1. (x1::int) - x0 = x1 + - 1 * x0" by force then have "x \ 2 ^ LENGTH('b) - 1" using f7 f6 f5 f4 by (metis uint_word_of_int wi_homs(2) word_arith_wis(8) word_of_int_2p) then have \uint x \ uint (2 ^ LENGTH('b) - (1 :: 'a word))\ by (simp add: word_le_def) then have \uint x \ 2 ^ LENGTH('b) - 1\ by (simp add: uint_word_ariths) (metis \1 \ 2 ^ LENGTH('b)\ \uint x \ uint (2 ^ LENGTH('b) - 1)\ linorder_not_less lt2p_lem uint_1 uint_minus_simple_alt uint_power_lower word_le_def zle_diff1_eq) then show ?thesis by (simp add: ucast_def word_ubin.eq_norm bintrunc_mod2p) qed lemma remdups_enum_upto: fixes s::"'a::len word" shows "remdups [s .e. e] = [s .e. e]" by simp lemma card_enum_upto: fixes s::"'a::len word" shows "card (set [s .e. e]) = Suc (unat e) - unat s" by (subst List.card_set) (simp add: remdups_enum_upto) lemma unat_mask: "unat (mask n :: 'a :: len word) = 2 ^ (min n (LENGTH('a))) - 1" apply (subst min.commute) apply (simp add: mask_def not_less min_def split: if_split_asm) apply (intro conjI impI) apply (simp add: unat_sub_if_size) apply (simp add: power_overflow word_size) apply (simp add: unat_sub_if_size) done lemma word_shiftr_lt: fixes w :: "'a::len word" shows "unat (w >> n) < (2 ^ (LENGTH('a) - n))" apply (subst shiftr_div_2n') by (metis nat_mod_lem nat_zero_less_power_iff power_mod_div word_unat.Rep_inverse word_unat.eq_norm zero_less_numeral) lemma complement_nth_w2p: shows "n' < LENGTH('a) \ (~~ (2 ^ n :: 'a::len word)) !! n' = (n' \ n)" by (fastforce simp: word_ops_nth_size word_size nth_w2p) lemma word_unat_and_lt: "unat x < n \ unat y < n \ unat (x && y) < n" by (meson le_less_trans word_and_le1 word_and_le2 word_le_nat_alt) lemma word_unat_mask_lt: "m \ size w \ unat ((w::'a::len word) && mask m) < 2 ^ m" by (rule word_unat_and_lt) (simp add: unat_mask word_size) lemma unat_shiftr_less_t2n: fixes x :: "'a :: len word" shows "unat x < 2 ^ (n + m) \ unat (x >> n) < 2 ^ m" by (simp add: shiftr_div_2n' power_add mult.commute td_gal_lt) lemma le_or_mask: "w \ w' \ w || mask x \ w' || mask x" by (metis neg_mask_add_mask add.commute le_word_or1 mask_2pm1 neg_mask_mono_le word_plus_mono_left) lemma le_shiftr1': "\ shiftr1 u \ shiftr1 v ; shiftr1 u \ shiftr1 v \ \ u \ v" apply (unfold word_le_def shiftr1_def word_ubin.eq_norm) apply (unfold bin_rest_trunc_i trans [OF bintrunc_bintrunc_l word_ubin.norm_Rep, unfolded word_ubin.norm_Rep, OF order_refl [THEN le_SucI]]) apply (case_tac "uint u" rule: bin_exhaust) apply (case_tac "uint v" rule: bin_exhaust) by (smt Bit_def bin_rest_BIT) lemma le_shiftr': - "\ u >> n \ v >> n ; u >> n \ v >> n \ \ (u::'a::len0 word) \ v" + "\ u >> n \ v >> n ; u >> n \ v >> n \ \ (u::'a::len word) \ v" apply (induct n; simp add: shiftr_def) apply (case_tac "(shiftr1 ^^ n) u = (shiftr1 ^^ n) v", simp) apply (fastforce dest: le_shiftr1') done lemma word_log2_nth_same: "w \ 0 \ w !! word_log2 w" unfolding word_log2_def using nth_length_takeWhile[where P=Not and xs="to_bl w"] apply (simp add: word_clz_def word_size to_bl_nth) apply (fastforce simp: linorder_not_less eq_zero_set_bl dest: takeWhile_take_has_property) done lemma word_log2_nth_not_set: "\ word_log2 w < i ; i < size w \ \ \ w !! i" unfolding word_log2_def word_clz_def using takeWhile_take_has_property_nth[where P=Not and xs="to_bl w" and n="size w - Suc i"] by (fastforce simp add: to_bl_nth word_size) lemma word_log2_highest: assumes a: "w !! i" shows "i \ word_log2 w" proof - from a have "i < size w" by - (rule test_bit_size) with a show ?thesis by - (rule ccontr, simp add: word_log2_nth_not_set) qed lemma word_log2_max: "word_log2 w < size w" unfolding word_log2_def word_clz_def by simp lemma word_clz_0[simp]: "word_clz (0::'a::len word) = LENGTH('a)" unfolding word_clz_def by (simp add: takeWhile_replicate) lemma word_clz_minus_one[simp]: "word_clz (-1::'a::len word) = 0" unfolding word_clz_def by (simp add: takeWhile_replicate) lemma word_add_no_overflow:"(x::'a::len word) < max_word \ x < x + 1" using less_x_plus_1 order_less_le by blast lemma lt_plus_1_le_word: fixes x :: "'a::len word" assumes bound:"n < unat (maxBound::'a word)" shows "x < 1 + of_nat n = (x \ of_nat n)" by (metis add.commute bound max_word_max word_Suc_leq word_not_le word_of_nat_less) lemma unat_ucast_up_simp: fixes x :: "'a::len word" assumes "LENGTH('a) \ LENGTH('b)" shows "unat (ucast x :: 'b::len word) = unat x" unfolding ucast_def unat_def apply (subst int_word_uint) apply (subst mod_pos_pos_trivial; simp?) apply (rule lt2p_lem) apply (simp add: assms) done lemma unat_ucast_less_no_overflow: "\n < 2 ^ LENGTH('a); unat f < n\ \ (f::('a::len) word) < of_nat n" by (erule (1) order_le_less_trans[OF _ of_nat_mono_maybe,rotated]) simp lemma unat_ucast_less_no_overflow_simp: "n < 2 ^ LENGTH('a) \ (unat f < n) = ((f::('a::len) word) < of_nat n)" using unat_less_helper unat_ucast_less_no_overflow by blast lemma unat_ucast_no_overflow_le: assumes no_overflow: "unat b < (2 :: nat) ^ LENGTH('a)" and upward_cast: "LENGTH('a) < LENGTH('b)" shows "(ucast (f::'a::len word) < (b :: 'b :: len word)) = (unat f < unat b)" proof - have LR: "ucast f < b \ unat f < unat b" apply (rule unat_less_helper) apply (simp add:ucast_nat_def) apply (rule_tac 'b1 = 'b in ucast_less_ucast[OF order.strict_implies_order, THEN iffD1]) apply (rule upward_cast) apply (simp add: ucast_ucast_mask less_mask_eq word_less_nat_alt unat_power_lower[OF upward_cast] no_overflow) done have RL: "unat f < unat b \ ucast f < b" proof- assume ineq: "unat f < unat b" have "ucast (f::'a::len word) < ((ucast (ucast b ::'a::len word)) :: 'b :: len word)" apply (simp add: ucast_less_ucast[OF order.strict_implies_order] upward_cast) apply (simp add: ucast_nat_def[symmetric]) apply (rule unat_ucast_less_no_overflow[OF no_overflow ineq]) done then show ?thesis apply (rule order_less_le_trans) apply (simp add:ucast_ucast_mask word_and_le2) done qed then show ?thesis by (simp add:RL LR iffI) qed lemmas ucast_up_mono = ucast_less_ucast[THEN iffD2] (* casting a long word to a shorter word and casting back to the long word is equal to the original long word -- if the word is small enough. 'l is the longer word. 's is the shorter word. *) lemma bl_cast_long_short_long_ingoreLeadingZero_generic: "\ length (dropWhile Not (to_bl w)) \ LENGTH('s); LENGTH('s) \ LENGTH('l) \ \ (of_bl :: _ \ 'l::len word) (to_bl ((of_bl::_ \ 's::len word) (to_bl w))) = w" by (rule word_uint_eqI) (simp add: uint_of_bl_is_bl_to_bin uint_of_bl_is_bl_to_bin_drop) (* Casting between longer and shorter word. 'l is the longer word. 's is the shorter word. For example: 'l::len word is 128 word (full ipv6 address) 's::len word is 16 word (address piece of ipv6 address in colon-text-representation) *) corollary ucast_short_ucast_long_ingoreLeadingZero: "\ length (dropWhile Not (to_bl w)) \ LENGTH('s); LENGTH('s) \ LENGTH('l) \ \ (ucast:: 's::len word \ 'l::len word) ((ucast:: 'l::len word \ 's::len word) w) = w" apply (subst Word.ucast_bl)+ apply (rule bl_cast_long_short_long_ingoreLeadingZero_generic; simp) done lemma length_drop_mask: fixes w::"'a::len word" shows "length (dropWhile Not (to_bl (w AND mask n))) \ n" proof - have "length (takeWhile Not (replicate n False @ ls)) = n + length (takeWhile Not ls)" for ls n by(subst takeWhile_append2) simp+ then show ?thesis unfolding bl_and_mask by (simp add: dropWhile_eq_drop) qed lemma minus_one_word: "(-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" by simp lemma mask_exceed: "n \ LENGTH('a) \ (x::'a::len word) && ~~ (mask n) = 0" by (simp add: and_not_mask shiftr_eq_0) lemma two_power_strict_part_mono: "strict_part_mono {..LENGTH('a) - 1} (\x. (2 :: 'a :: len word) ^ x)" proof - { fix n have "n < LENGTH('a) \ strict_part_mono {..n} (\x. (2 :: 'a :: len word) ^ x)" proof (induct n) case 0 then show ?case by simp next case (Suc n) from Suc.prems have "2 ^ n < (2 :: 'a :: len word) ^ Suc n" using power_strict_increasing unat_power_lower word_less_nat_alt by fastforce with Suc show ?case by (subst strict_part_mono_by_steps) simp qed } then show ?thesis by simp qed lemma word_shift_by_2: "x * 4 = (x::'a::len word) << 2" by (simp add: shiftl_t2n) lemma le_2p_upper_bits: "\ (p::'a::len word) \ 2^n - 1; n < LENGTH('a) \ \ \n'\n. n' < LENGTH('a) \ \ p !! n'" by (subst upper_bits_unset_is_l2p; simp) lemma le2p_bits_unset: "p \ 2 ^ n - 1 \ \n'\n. n' < LENGTH('a) \ \ (p::'a::len word) !! n'" using upper_bits_unset_is_l2p [where p=p] by (cases "n < LENGTH('a)") auto lemma ucast_less_shiftl_helper: "\ LENGTH('b) + 2 < LENGTH('a); 2 ^ (LENGTH('b) + 2) \ n\ \ (ucast (x :: 'b::len word) << 2) < (n :: 'a::len word)" apply (erule order_less_le_trans[rotated]) using ucast_less[where x=x and 'a='a] apply (simp only: shiftl_t2n field_simps) apply (rule word_less_power_trans2; simp) done lemma word_power_nonzero: "\ (x :: 'a::len word) < 2 ^ (LENGTH('a) - n); n < LENGTH('a); x \ 0 \ \ x * 2 ^ n \ 0" by (metis and_mask_eq_iff_shiftr_0 less_mask_eq p2_gt_0 semiring_normalization_rules(7) shiftl_shiftr_id shiftl_t2n) lemma less_1_helper: "n \ m \ (n - 1 :: int) < m" by arith lemma div_power_helper: "\ x \ y; y < LENGTH('a) \ \ (2 ^ y - 1) div (2 ^ x :: 'a::len word) = 2 ^ (y - x) - 1" apply (rule word_uint.Rep_eqD) apply (simp only: uint_word_ariths uint_div uint_power_lower) apply (subst mod_pos_pos_trivial, fastforce, fastforce)+ apply (subst mod_pos_pos_trivial) apply (simp add: le_diff_eq uint_2p_alt) apply (rule less_1_helper) apply (rule power_increasing; simp) apply (subst mod_pos_pos_trivial) apply (simp add: uint_2p_alt) apply (rule less_1_helper) apply (rule power_increasing; simp) apply (subst int_div_sub_1; simp add: uint_2p_alt) apply (subst power_0[symmetric]) apply (simp add: uint_2p_alt le_imp_power_dvd power_sub_int) done lemma word_add_power_off: fixes a :: "'a :: len word" assumes ak: "a < k" and kw: "k < 2 ^ (LENGTH('a) - m)" and mw: "m < LENGTH('a)" and off: "off < 2 ^ m" shows "(a * 2 ^ m) + off < k * 2 ^ m" proof (cases "m = 0") case True then show ?thesis using off ak by simp next case False from ak have ak1: "a + 1 \ k" by (rule inc_le) then have "(a + 1) * 2 ^ m \ 0" apply - apply (rule word_power_nonzero) apply (erule order_le_less_trans [OF _ kw]) apply (rule mw) apply (rule less_is_non_zero_p1 [OF ak]) done then have "(a * 2 ^ m) + off < ((a + 1) * 2 ^ m)" using kw mw apply - apply (simp add: distrib_right) apply (rule word_plus_strict_mono_right [OF off]) apply (rule is_aligned_no_overflow'') apply (rule is_aligned_mult_triv2) apply assumption done also have "\ \ k * 2 ^ m" using ak1 mw kw False apply - apply (erule word_mult_le_mono1) apply (simp add: p2_gt_0) apply (simp add: word_less_nat_alt) apply (rule nat_less_power_trans2[simplified]) apply (simp add: word_less_nat_alt) apply simp done finally show ?thesis . qed lemma offset_not_aligned: "\ is_aligned (p::'a::len word) n; i > 0; i < 2 ^ n; n < LENGTH('a)\ \ \ is_aligned (p + of_nat i) n" apply (erule is_aligned_add_not_aligned) unfolding is_aligned_def by (metis le_unat_uoi nat_dvd_not_less order_less_imp_le unat_power_lower) lemma length_upto_enum_one: fixes x :: "'a :: len word" assumes lt1: "x < y" and lt2: "z < y" and lt3: "x \ z" shows "[x , y .e. z] = [x]" unfolding upto_enum_step_def proof (subst upto_enum_red, subst if_not_P [OF leD [OF lt3]], clarsimp, rule conjI) show "unat ((z - x) div (y - x)) = 0" proof (subst unat_div, rule div_less) have syx: "unat (y - x) = unat y - unat x" by (rule unat_sub [OF order_less_imp_le]) fact moreover have "unat (z - x) = unat z - unat x" by (rule unat_sub) fact ultimately show "unat (z - x) < unat (y - x)" using lt2 lt3 unat_mono word_less_minus_mono_left by blast qed then show "(z - x) div (y - x) * (y - x) = 0" by (metis mult_zero_left unat_0 word_unat.Rep_eqD) qed lemma max_word_mask: "(max_word :: 'a::len word) = mask LENGTH('a)" unfolding mask_def by simp lemmas mask_len_max = max_word_mask[symmetric] lemma is_aligned_alignUp[simp]: "is_aligned (alignUp p n) n" by (simp add: alignUp_def complement_def is_aligned_mask mask_def word_bw_assocs) lemma alignUp_le[simp]: "alignUp p n \ p + 2 ^ n - 1" unfolding alignUp_def by (rule word_and_le2) lemma complement_mask: "complement (2 ^ n - 1) = ~~ (mask n)" unfolding complement_def mask_def by simp lemma alignUp_idem: fixes a :: "'a::len word" assumes "is_aligned a n" "n < LENGTH('a)" shows "alignUp a n = a" using assms unfolding alignUp_def by (metis complement_mask is_aligned_add_helper p_assoc_help power_2_ge_iff) lemma alignUp_not_aligned_eq: fixes a :: "'a :: len word" assumes al: "\ is_aligned a n" and sz: "n < LENGTH('a)" shows "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" proof - have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) fact+ then have um: "unat (a mod 2 ^ n - 1) div 2 ^ n = 0" using sz by (meson Euclidean_Division.div_eq_0_iff le_m1_iff_lt measure_unat order_less_trans unat_less_power word_less_sub_le word_mod_less_divisor) have "a + 2 ^ n - 1 = (a div 2 ^ n) * 2 ^ n + (a mod 2 ^ n) + 2 ^ n - 1" by (simp add: word_mod_div_equality) also have "\ = (a mod 2 ^ n - 1) + (a div 2 ^ n + 1) * 2 ^ n" by (simp add: field_simps) finally show "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" using sz unfolding alignUp_def apply (subst complement_mask) apply (erule ssubst) apply (subst neg_mask_is_div) apply (simp add: word_arith_nat_div) apply (subst unat_word_ariths(1) unat_word_ariths(2))+ apply (subst uno_simps) apply (subst unat_1) apply (subst mod_add_right_eq) apply simp apply (subst power_mod_div) apply (subst div_mult_self1) apply simp apply (subst um) apply simp apply (subst mod_mod_power) apply simp apply (subst word_unat_power, subst Abs_fnat_hom_mult) apply (subst mult_mod_left) apply (subst power_add [symmetric]) apply simp apply (subst Abs_fnat_hom_1) apply (subst Abs_fnat_hom_add) apply (subst word_unat_power, subst Abs_fnat_hom_mult) apply (subst word_unat.Rep_inverse[symmetric], subst Abs_fnat_hom_mult) apply simp done qed lemma alignUp_ge: fixes a :: "'a :: len word" assumes sz: "n < LENGTH('a)" and nowrap: "alignUp a n \ 0" shows "a \ alignUp a n" proof (cases "is_aligned a n") case True then show ?thesis using sz by (subst alignUp_idem, simp_all) next case False have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz by (metis shiftr_div_2n' word_shiftr_lt) have"2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using sz by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right lt0 nat_le_power_trans nat_less_le) moreover have "2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using nowrap sz apply - apply (erule contrapos_nn) apply (subst alignUp_not_aligned_eq [OF False sz]) apply (subst unat_arith_simps) apply (subst unat_word_ariths) apply (subst unat_word_ariths) apply simp apply (subst mult_mod_left) apply (simp add: unat_div field_simps power_add[symmetric] mod_mod_power min.absorb2) done ultimately have lt: "2 ^ n * (unat a div 2 ^ n + 1) < 2 ^ LENGTH('a)" by simp have "a = a div 2 ^ n * 2 ^ n + a mod 2 ^ n" by (rule word_mod_div_equality [symmetric]) also have "\ < (a div 2 ^ n + 1) * 2 ^ n" using sz lt apply (simp add: field_simps) apply (rule word_add_less_mono1) apply (rule word_mod_less_divisor) apply (simp add: word_less_nat_alt) apply (subst unat_word_ariths) apply (simp add: unat_div) done also have "\ = alignUp a n" by (rule alignUp_not_aligned_eq [symmetric]) fact+ finally show ?thesis by (rule order_less_imp_le) qed lemma alignUp_le_greater_al: fixes x :: "'a :: len word" assumes le: "a \ x" and sz: "n < LENGTH('a)" and al: "is_aligned x n" shows "alignUp a n \ x" proof (cases "is_aligned a n") case True then show ?thesis using sz le by (simp add: alignUp_idem) next case False then have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) from al obtain k where xk: "x = 2 ^ n * of_nat k" and kv: "k < 2 ^ (LENGTH('a) - n)" by (auto elim!: is_alignedE) then have kn: "unat (of_nat k :: 'a word) * unat ((2::'a word) ^ n) < 2 ^ LENGTH('a)" using sz apply (subst unat_of_nat_eq) apply (erule order_less_le_trans) apply simp apply (subst mult.commute) apply simp apply (rule nat_less_power_trans) apply simp apply simp done have au: "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" by (rule alignUp_not_aligned_eq) fact+ also have "\ \ of_nat k * 2 ^ n" proof (rule word_mult_le_mono1 [OF inc_le _ kn]) show "a div 2 ^ n < of_nat k" using kv xk le sz anz by (simp add: alignUp_div_helper) show "(0:: 'a word) < 2 ^ n" using sz by (simp add: p2_gt_0 sz) qed finally show ?thesis using xk by (simp add: field_simps) qed lemma alignUp_is_aligned_nz: fixes a :: "'a :: len word" assumes al: "is_aligned x n" and sz: "n < LENGTH('a)" and ax: "a \ x" and az: "a \ 0" shows "alignUp (a::'a :: len word) n \ 0" proof (cases "is_aligned a n") case True then have "alignUp a n = a" using sz by (simp add: alignUp_idem) then show ?thesis using az by simp next case False then have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) { assume asm: "alignUp a n = 0" have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz by (metis shiftr_div_2n' word_shiftr_lt) have leq: "2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using sz by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right lt0 nat_le_power_trans order_less_imp_le) from al obtain k where kv: "k < 2 ^ (LENGTH('a) - n)" and xk: "x = 2 ^ n * of_nat k" by (auto elim!: is_alignedE) then have "a div 2 ^ n < of_nat k" using ax sz anz by (rule alignUp_div_helper) then have r: "unat a div 2 ^ n < k" using sz by (metis unat_div unat_less_helper unat_power_lower) have "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" by (rule alignUp_not_aligned_eq) fact+ then have "\ = 0" using asm by simp then have "2 ^ LENGTH('a) dvd 2 ^ n * (unat a div 2 ^ n + 1)" using sz by (simp add: unat_arith_simps ac_simps) (simp add: unat_word_ariths mod_simps mod_eq_0_iff_dvd) with leq have "2 ^ n * (unat a div 2 ^ n + 1) = 2 ^ LENGTH('a)" by (force elim!: le_SucE) then have "unat a div 2 ^ n = 2 ^ LENGTH('a) div 2 ^ n - 1" by (metis (no_types, hide_lams) Groups.add_ac(2) add.right_neutral add_diff_cancel_left' div_le_dividend div_mult_self4 gr_implies_not0 le_neq_implies_less power_eq_0_iff unat_def zero_neq_numeral) then have "unat a div 2 ^ n = 2 ^ (LENGTH('a) - n) - 1" using sz by (simp add: power_sub) then have "2 ^ (LENGTH('a) - n) - 1 < k" using r by simp then have False using kv by simp } then show ?thesis by clarsimp qed lemma alignUp_ar_helper: fixes a :: "'a :: len word" assumes al: "is_aligned x n" and sz: "n < LENGTH('a)" and sub: "{x..x + 2 ^ n - 1} \ {a..b}" and anz: "a \ 0" shows "a \ alignUp a n \ alignUp a n + 2 ^ n - 1 \ b" proof from al have xl: "x \ x + 2 ^ n - 1" by (simp add: is_aligned_no_overflow) from xl sub have ax: "a \ x" by (clarsimp elim!: range_subset_lower [where x = x]) show "a \ alignUp a n" proof (rule alignUp_ge) show "alignUp a n \ 0" using al sz ax anz by (rule alignUp_is_aligned_nz) qed fact+ show "alignUp a n + 2 ^ n - 1 \ b" proof (rule order_trans) from xl show tp: "x + 2 ^ n - 1 \ b" using sub by (clarsimp elim!: range_subset_upper [where x = x]) from ax have "alignUp a n \ x" by (rule alignUp_le_greater_al) fact+ then have "alignUp a n + (2 ^ n - 1) \ x + (2 ^ n - 1)" using xl al is_aligned_no_overflow' olen_add_eqv word_plus_mcs_3 by blast then show "alignUp a n + 2 ^ n - 1 \ x + 2 ^ n - 1" by (simp add: field_simps) qed qed lemma alignUp_def2: "alignUp a sz = a + 2 ^ sz - 1 && ~~ (mask sz)" unfolding alignUp_def[unfolded complement_def] by (simp add:mask_def[symmetric,unfolded shiftl_t2n,simplified]) lemma mask_out_first_mask_some: "\ x && ~~ (mask n) = y; n \ m \ \ x && ~~ (mask m) = y && ~~ (mask m)" by word_eqI_solve lemma gap_between_aligned: "\a < (b :: 'a ::len word); is_aligned a n; is_aligned b n; n < LENGTH('a) \ \ a + (2^n - 1) < b" by (simp add: aligned_add_offset_less) lemma mask_out_add_aligned: assumes al: "is_aligned p n" shows "p + (q && ~~ (mask n)) = (p + q) && ~~ (mask n)" using mask_add_aligned [OF al] by (simp add: mask_out_sub_mask) lemma alignUp_def3: "alignUp a sz = 2^ sz + (a - 1 && ~~ (mask sz))" by (simp add: alignUp_def2 is_aligned_triv field_simps mask_out_add_aligned) lemma alignUp_plus: "is_aligned w us \ alignUp (w + a) us = w + alignUp a us" by (clarsimp simp: alignUp_def2 mask_out_add_aligned field_simps) lemma mask_lower_twice: "n \ m \ (x && ~~ (mask n)) && ~~ (mask m) = x && ~~ (mask m)" by word_eqI_solve lemma mask_lower_twice2: "(a && ~~ (mask n)) && ~~ (mask m) = a && ~~ (mask (max n m))" by word_eqI_solve lemma ucast_and_neg_mask: "ucast (x && ~~ (mask n)) = ucast x && ~~ (mask n)" by word_eqI_solve lemma ucast_and_mask: "ucast (x && mask n) = ucast x && mask n" by word_eqI_solve lemma ucast_mask_drop: "LENGTH('a :: len) \ n \ (ucast (x && mask n) :: 'a word) = ucast x" by word_eqI lemma alignUp_distance: "alignUp (q :: 'a :: len word) sz - q \ mask sz" by (metis (no_types) add.commute add_diff_cancel_left alignUp_def2 diff_add_cancel mask_2pm1 subtract_mask(2) word_and_le1 word_sub_le_iff) lemma is_aligned_diff_neg_mask: "is_aligned p sz \ (p - q && ~~ (mask sz)) = (p - ((alignUp q sz) && ~~ (mask sz)))" apply (clarsimp simp only:word_and_le2 diff_conv_add_uminus) apply (subst mask_out_add_aligned[symmetric]; simp) apply (rule sum_to_zero) apply (subst add.commute) by (simp add: alignUp_distance and_mask_0_iff_le_mask is_aligned_neg_mask_eq mask_out_add_aligned) lemma map_bits_rev_to_bl: "map ((!!) x) [0.. LENGTH('a) \ x = ucast y \ ucast x = y" for x :: "'a::len word" and y :: "'b::len word" by (simp add: is_down ucast_id ucast_ucast_a) lemma le_ucast_ucast_le: "x \ ucast y \ ucast x \ y" for x :: "'a::len word" and y :: "'b::len word" by (smt le_unat_uoi linorder_not_less order_less_imp_le ucast_nat_def unat_arith_simps(1)) lemma less_ucast_ucast_less: "LENGTH('b) \ LENGTH('a) \ x < ucast y \ ucast x < y" for x :: "'a::len word" and y :: "'b::len word" by (metis ucast_nat_def unat_mono unat_ucast_up_simp word_of_nat_less) lemma ucast_le_ucast: "LENGTH('a) \ LENGTH('b) \ (ucast x \ (ucast y::'b::len word)) = (x \ y)" for x :: "'a::len word" by (simp add: unat_arith_simps(1) unat_ucast_up_simp) lemmas ucast_up_mono_le = ucast_le_ucast[THEN iffD2] lemma ucast_le_ucast_eq: fixes x y :: "'a::len word" assumes x: "x < 2 ^ n" assumes y: "y < 2 ^ n" assumes n: "n = LENGTH('b::len)" shows "(UCAST('a \ 'b) x \ UCAST('a \ 'b) y) = (x \ y)" apply (rule iffI) apply (cases "LENGTH('b) < LENGTH('a)") apply (subst less_mask_eq[OF x, symmetric]) apply (subst less_mask_eq[OF y, symmetric]) apply (unfold n) apply (subst ucast_ucast_mask[symmetric])+ apply (simp add: ucast_le_ucast)+ apply (erule ucast_mono_le[OF _ y[unfolded n]]) done lemma word_le_not_less: "((b::'a::len word) \ a) = (\(a < b))" by fastforce lemma ucast_or_distrib: fixes x :: "'a::len word" fixes y :: "'a::len word" shows "(ucast (x || y) :: ('b::len) word) = ucast x || ucast y" unfolding ucast_def Word.bitOR_word.abs_eq uint_or by blast lemma shiftr_less: "(w::'a::len word) < k \ w >> n < k" by (metis div_le_dividend le_less_trans shiftr_div_2n' unat_arith_simps(2)) lemma word_and_notzeroD: "w && w' \ 0 \ w \ 0 \ w' \ 0" by auto lemma word_clz_max: "word_clz w \ size (w::'a::len word)" unfolding word_clz_def by (metis length_takeWhile_le word_size_bl) lemma word_clz_nonzero_max: fixes w :: "'a::len word" assumes nz: "w \ 0" shows "word_clz w < size (w::'a::len word)" proof - { assume a: "word_clz w = size (w::'a::len word)" hence "length (takeWhile Not (to_bl w)) = length (to_bl w)" by (simp add: word_clz_def word_size) hence allj: "\j\set(to_bl w). \ j" by (metis a length_takeWhile_less less_irrefl_nat word_clz_def) hence "to_bl w = replicate (length (to_bl w)) False" by (fastforce intro!: list_of_false) hence "w = 0" by (metis to_bl_0 word_bl.Rep_eqD word_bl_Rep') with nz have False by simp } thus ?thesis using word_clz_max by (fastforce intro: le_neq_trans) qed lemma unat_add_lem': "(unat x + unat y < 2 ^ LENGTH('a)) \ (unat (x + y :: 'a :: len word) = unat x + unat y)" by (subst unat_add_lem[symmetric], assumption) lemma from_bool_eq_if': "((if P then 1 else 0) = from_bool Q) = (P = Q)" by (simp add: case_bool_If from_bool_def split: if_split) lemma word_exists_nth: "(w::'a::len word) \ 0 \ \i. w !! i" using word_log2_nth_same by blast lemma shiftr_le_0: "unat (w::'a::len word) < 2 ^ n \ w >> n = (0::'a::len word)" by (rule word_unat.Rep_eqD) (simp add: shiftr_div_2n') lemma of_nat_shiftl: "(of_nat x << n) = (of_nat (x * 2 ^ n) :: ('a::len) word)" proof - have "(of_nat x::'a word) << n = of_nat (2 ^ n) * of_nat x" using shiftl_t2n by (metis word_unat_power) thus ?thesis by simp qed lemma shiftl_1_not_0: "n < LENGTH('a) \ (1::'a::len word) << n \ 0" by (simp add: shiftl_t2n) lemma max_word_not_0 [simp]: "- 1 \ (0 :: 'a::len word)" by simp lemma ucast_zero_is_aligned: "UCAST('a::len \ 'b::len) w = 0 \ n \ LENGTH('b) \ is_aligned w n" by (clarsimp simp: is_aligned_mask word_eq_iff word_size nth_ucast) lemma unat_ucast_eq_unat_and_mask: "unat (UCAST('b::len \ 'a::len) w) = unat (w && mask LENGTH('a))" proof - have "unat (UCAST('b \ 'a) w) = unat (UCAST('a \ 'b) (UCAST('b \ 'a) w))" by (cases "LENGTH('a) < LENGTH('b)"; simp add: is_down ucast_ucast_a unat_ucast_up_simp) thus ?thesis using ucast_ucast_mask by simp qed lemma unat_max_word_pos[simp]: "0 < unat (- 1 :: 'a::len word)" using unat_gt_0 [of \- 1 :: 'a::len word\] by simp (* Miscellaneous conditional injectivity rules. *) lemma mult_pow2_inj: assumes ws: "m + n \ LENGTH('a)" assumes le: "x \ mask m" "y \ mask m" assumes eq: "x * 2^n = y * (2^n::'a::len word)" shows "x = y" proof (cases n) case 0 thus ?thesis using eq by simp next case (Suc n') have m_lt: "m < LENGTH('a)" using Suc ws by simp have xylt: "x < 2^m" "y < 2^m" using le m_lt unfolding mask_2pm1 by auto have lenm: "n \ LENGTH('a) - m" using ws by simp show ?thesis using eq xylt apply (fold shiftl_t2n[where n=n, simplified mult.commute]) apply (simp only: word_bl.Rep_inject[symmetric] bl_shiftl) apply (erule ssubst[OF less_is_drop_replicate])+ apply (clarsimp elim!: drop_eq_mono[OF lenm]) done qed lemma word_of_nat_inj: assumes bounded: "x < 2 ^ LENGTH('a)" "y < 2 ^ LENGTH('a)" assumes of_nats: "of_nat x = (of_nat y :: 'a::len word)" shows "x = y" by (rule contrapos_pp[OF of_nats]; cases "x < y"; cases "y < x") (auto dest: bounded[THEN of_nat_mono_maybe]) (* Sign extension from bit n. *) lemma sign_extend_bitwise_if: "i < size w \ sign_extend e w !! i \ (if i < e then w !! i else w !! e)" by (simp add: sign_extend_def neg_mask_test_bit word_size) lemma sign_extend_bitwise_disj: "i < size w \ sign_extend e w !! i \ i \ e \ w !! i \ e \ i \ w !! e" by (auto simp: sign_extend_bitwise_if) lemma sign_extend_bitwise_cases: "i < size w \ sign_extend e w !! i \ (i \ e \ w !! i) \ (e \ i \ w !! e)" by (auto simp: sign_extend_bitwise_if) lemmas sign_extend_bitwise_if'[word_eqI_simps] = sign_extend_bitwise_if[simplified word_size] lemmas sign_extend_bitwise_disj' = sign_extend_bitwise_disj[simplified word_size] lemmas sign_extend_bitwise_cases' = sign_extend_bitwise_cases[simplified word_size] (* Often, it is easier to reason about an operation which does not overwrite the bit which determines which mask operation to apply. *) lemma sign_extend_def': "sign_extend n w = (if w !! n then w || ~~ (mask (Suc n)) else w && mask (Suc n))" by word_eqI (auto dest: less_antisym) lemma sign_extended_sign_extend: "sign_extended n (sign_extend n w)" by (clarsimp simp: sign_extended_def word_size sign_extend_bitwise_if') lemma sign_extended_iff_sign_extend: "sign_extended n w \ sign_extend n w = w" apply (rule iffI) apply (word_eqI, rename_tac i) apply (case_tac "n < i"; simp add: sign_extended_def word_size) apply (erule subst, rule sign_extended_sign_extend) done lemma sign_extended_weaken: "sign_extended n w \ n \ m \ sign_extended m w" unfolding sign_extended_def by (cases "n < m") auto lemma sign_extend_sign_extend_eq: "sign_extend m (sign_extend n w) = sign_extend (min m n) w" by word_eqI lemma sign_extended_high_bits: "\ sign_extended e p; j < size p; e \ i; i < j \ \ p !! i = p !! j" by (drule (1) sign_extended_weaken; simp add: sign_extended_def) lemma sign_extend_eq: "w && mask (Suc n) = v && mask (Suc n) \ sign_extend n w = sign_extend n v" by word_eqI_solve lemma sign_extended_add: assumes p: "is_aligned p n" assumes f: "f < 2 ^ n" assumes e: "n \ e" assumes "sign_extended e p" shows "sign_extended e (p + f)" proof (cases "e < size p") case True note and_or = is_aligned_add_or[OF p f] have "\ f !! e" using True e less_2p_is_upper_bits_unset[THEN iffD1, OF f] by (fastforce simp: word_size) hence i: "(p + f) !! e = p !! e" by (simp add: and_or) have fm: "f && mask e = f" by (fastforce intro: subst[where P="\f. f && mask e = f", OF less_mask_eq[OF f]] simp: mask_twice e) show ?thesis using assms apply (simp add: sign_extended_iff_sign_extend sign_extend_def i) apply (simp add: and_or word_bw_comms[of p f]) apply (clarsimp simp: word_ao_dist fm word_bw_assocs split: if_splits) done next case False thus ?thesis by (simp add: sign_extended_def word_size) qed lemma sign_extended_neq_mask: "\sign_extended n ptr; m \ n\ \ sign_extended n (ptr && ~~ (mask m))" by (fastforce simp: sign_extended_def word_size neg_mask_test_bit) (* Uints *) lemma uints_mono_iff: "uints l \ uints m \ l \ m" using power_increasing_iff[of "2::int" l m] apply (auto simp: uints_num subset_iff simp del: power_increasing_iff) by (meson less_irrefl not_less zle2p) lemmas uints_monoI = uints_mono_iff[THEN iffD2] lemma Bit_in_uints_Suc: "w BIT c \ uints (Suc m)" if "w \ uints m" using that by (auto simp: uints_num Bit_def) lemma Bit_in_uintsI: "w BIT c \ uints m" if "w \ uints (m - 1)" "m > 0" using Bit_in_uints_Suc[OF that(1)] that(2) by auto lemma bin_cat_in_uintsI: "bin_cat a n b \ uints m" if "a \ uints l" "m \ l + n" using that proof (induction n arbitrary: b m) case 0 then have "uints l \ uints m" by (intro uints_monoI) auto then show ?case using 0 by auto next case (Suc n) then show ?case by (auto intro!: Bit_in_uintsI) qed lemma bin_cat_cong: "bin_cat a n b = bin_cat c m d" if "n = m" "a = c" "bintrunc m b = bintrunc m d" using that(3) unfolding that(1,2) by (simp add: bin_cat_eq_push_bit_add_take_bit) lemma bin_cat_eqD1: "bin_cat a n b = bin_cat c n d \ a = c" by (induction n arbitrary: b d) auto lemma bin_cat_eqD2: "bin_cat a n b = bin_cat c n d \ bintrunc n b = bintrunc n d" by (induction n arbitrary: b d) (simp_all add: take_bit_Suc) lemma bin_cat_inj: "(bin_cat a n b) = bin_cat c n d \ a = c \ bintrunc n b = bintrunc n d" by (auto intro: bin_cat_cong bin_cat_eqD1 bin_cat_eqD2) lemma word_of_int_bin_cat_eq_iff: - "(word_of_int (bin_cat (uint a) LENGTH('b) (uint b))::'c::len0 word) = + "(word_of_int (bin_cat (uint a) LENGTH('b) (uint b))::'c::len word) = word_of_int (bin_cat (uint c) LENGTH('b) (uint d)) \ b = d \ a = c" if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" - for a::"'a::len0 word" and b::"'b::len0 word" + for a::"'a::len word" and b::"'b::len word" by (subst word_uint.Abs_inject) (auto simp: bin_cat_inj intro!: that bin_cat_in_uintsI) -lemma word_cat_inj: "(word_cat a b::'c::len0 word) = word_cat c d \ a = c \ b = d" +lemma word_cat_inj: "(word_cat a b::'c::len word) = word_cat c d \ a = c \ b = d" if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" - for a::"'a::len0 word" and b::"'b::len0 word" + for a::"'a::len word" and b::"'b::len word" by (auto simp: word_cat_def word_uint.Abs_inject word_of_int_bin_cat_eq_iff that) lemma p2_eq_1: "2 ^ n = (1::'a::len word) \ n = 0" proof - have "2 ^ n = (1::'a word) \ n = 0" by (metis One_nat_def not_less one_less_numeral_iff p2_eq_0 p2_gt_0 power_0 power_0 power_inject_exp semiring_norm(76) unat_power_lower zero_neq_one) then show ?thesis by auto qed (* usually: x,y = (len_of TYPE ('a)) *) lemma bitmagic_zeroLast_leq_or1Last: "(a::('a::len) word) AND (mask len << x - len) \ a OR mask (y - len)" by (meson le_word_or2 order_trans word_and_le2) lemma zero_base_lsb_imp_set_eq_as_bit_operation: fixes base ::"'a::len word" assumes valid_prefix: "mask (LENGTH('a) - len) AND base = 0" shows "(base = NOT (mask (LENGTH('a) - len)) AND a) \ (a \ {base .. base OR mask (LENGTH('a) - len)})" proof have helper3: "x OR y = x OR y AND NOT x" for x y ::"'a::len word" by (simp add: word_oa_dist2) from assms show "base = NOT (mask (LENGTH('a) - len)) AND a \ a \ {base..base OR mask (LENGTH('a) - len)}" apply(simp add: word_and_le1) apply(metis helper3 le_word_or2 word_bw_comms(1) word_bw_comms(2)) done next assume "a \ {base..base OR mask (LENGTH('a) - len)}" hence a: "base \ a \ a \ base OR mask (LENGTH('a) - len)" by simp show "base = NOT (mask (LENGTH('a) - len)) AND a" proof - have f2: "\x\<^sub>0. base AND NOT (mask x\<^sub>0) \ a AND NOT (mask x\<^sub>0)" using a neg_mask_mono_le by blast have f3: "\x\<^sub>0. a AND NOT (mask x\<^sub>0) \ (base OR mask (LENGTH('a) - len)) AND NOT (mask x\<^sub>0)" using a neg_mask_mono_le by blast have f4: "base = base AND NOT (mask (LENGTH('a) - len))" using valid_prefix by (metis mask_eq_0_eq_x word_bw_comms(1)) hence f5: "\x\<^sub>6. (base OR x\<^sub>6) AND NOT (mask (LENGTH('a) - len)) = base OR x\<^sub>6 AND NOT (mask (LENGTH('a) - len))" using word_ao_dist by (metis) have f6: "\x\<^sub>2 x\<^sub>3. a AND NOT (mask x\<^sub>2) \ x\<^sub>3 \ \ (base OR mask (LENGTH('a) - len)) AND NOT (mask x\<^sub>2) \ x\<^sub>3" using f3 dual_order.trans by auto have "base = (base OR mask (LENGTH('a) - len)) AND NOT (mask (LENGTH('a) - len))" using f5 by auto hence "base = a AND NOT (mask (LENGTH('a) - len))" using f2 f4 f6 by (metis eq_iff) thus "base = NOT (mask (LENGTH('a) - len)) AND a" by (metis word_bw_comms(1)) qed qed lemma unat_minus_one_word: "unat (-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" by (subst minus_one_word) (subst unat_sub_if', clarsimp) lemma of_nat_eq_signed_scast: "(of_nat x = (y :: ('a::len) signed word)) = (of_nat x = (scast y :: 'a word))" by (metis scast_of_nat scast_scast_id(2)) lemma word_ctz_le: "word_ctz (w :: ('a::len word)) \ LENGTH('a)" apply (clarsimp simp: word_ctz_def) apply (rule nat_le_Suc_less_imp[where y="LENGTH('a) + 1" , simplified]) apply (rule order_le_less_trans[OF List.length_takeWhile_le]) apply simp done lemma word_ctz_less: "w \ 0 \ word_ctz (w :: ('a::len word)) < LENGTH('a)" apply (clarsimp simp: word_ctz_def eq_zero_set_bl) apply (rule order_less_le_trans[OF length_takeWhile_less]) apply fastforce+ done lemma word_ctz_not_minus_1: "1 < LENGTH('a) \ of_nat (word_ctz (w :: 'a :: len word)) \ (- 1 :: 'a::len word)" by (metis (mono_tags) One_nat_def add.right_neutral add_Suc_right le_diff_conv le_less_trans n_less_equal_power_2 not_le suc_le_pow_2 unat_minus_one_word unat_of_nat_len word_ctz_le) lemma word_aligned_add_no_wrap_bounded: "\ w + 2^n \ x; w + 2^n \ 0; is_aligned w n \ \ (w::'a::len word) < x" by (blast dest: is_aligned_no_overflow le_less_trans word_leq_le_minus_one) lemma mask_Suc: "mask (Suc n) = 2^n + mask n" by (simp add: mask_def) lemma is_aligned_no_overflow_mask: "is_aligned x n \ x \ x + mask n" by (simp add: mask_def) (erule is_aligned_no_overflow') lemma is_aligned_mask_offset_unat: fixes off :: "('a::len) word" and x :: "'a word" assumes al: "is_aligned x sz" and offv: "off \ mask sz" shows "unat x + unat off < 2 ^ LENGTH('a)" proof cases assume szv: "sz < LENGTH('a)" from al obtain k where xv: "x = 2 ^ sz * (of_nat k)" and kl: "k < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) from offv szv have offv': "unat off < 2 ^ sz" by (simp add: mask_2pm1 unat_less_power) show ?thesis using szv using al is_aligned_no_wrap''' offv' by blast next assume "\ sz < LENGTH('a)" with al have "x = 0" by - word_eqI thus ?thesis by simp qed lemma of_bl_max: "(of_bl xs :: 'a::len word) \ mask (length xs)" apply (induct xs) apply simp apply (simp add: of_bl_Cons mask_Suc) apply (rule conjI; clarsimp) apply (erule word_plus_mono_right) apply (rule is_aligned_no_overflow_mask) apply (rule is_aligned_triv) apply (simp add: word_le_nat_alt) apply (subst unat_add_lem') apply (rule is_aligned_mask_offset_unat) apply (rule is_aligned_triv) apply (simp add: mask_def) apply simp done lemma mask_over_length: "LENGTH('a) \ n \ mask n = (-1::'a::len word)" by (simp add: mask_def) lemma is_aligned_over_length: "\ is_aligned p n; LENGTH('a) \ n \ \ (p::'a::len word) = 0" by (simp add: is_aligned_mask mask_over_length) lemma Suc_2p_unat_mask: "n < LENGTH('a) \ Suc (2 ^ n * k + unat (mask n :: 'a::len word)) = 2 ^ n * (k+1)" by (simp add: unat_mask) lemma is_aligned_add_step_le: "\ is_aligned (a::'a::len word) n; is_aligned b n; a < b; b \ a + mask n \ \ False" apply (simp flip: not_le) apply (erule notE) apply (cases "LENGTH('a) \ n") apply (drule (1) is_aligned_over_length)+ apply (drule mask_over_length) apply clarsimp apply (clarsimp simp: word_le_nat_alt not_less) apply (subst (asm) unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask) apply (clarsimp simp: is_aligned_def dvd_def word_le_nat_alt) apply (drule le_imp_less_Suc) apply (simp add: Suc_2p_unat_mask) by (metis Groups.mult_ac(2) Suc_leI linorder_not_less mult_le_mono order_refl times_nat.simps(2)) lemma power_2_mult_step_le: "\n' \ n; 2 ^ n' * k' < 2 ^ n * k\ \ 2 ^ n' * (k' + 1) \ 2 ^ n * (k::nat)" apply (cases "n'=n", simp) apply (metis Suc_leI le_refl mult_Suc_right mult_le_mono semiring_normalization_rules(7)) apply (drule (1) le_neq_trans) apply clarsimp apply (subgoal_tac "\m. n = n' + m") prefer 2 apply (simp add: le_Suc_ex) apply (clarsimp simp: power_add) by (metis Suc_leI mult.assoc mult_Suc_right nat_mult_le_cancel_disj) lemma aligned_mask_step: "\ n' \ n; p' \ p + mask n; is_aligned p n; is_aligned p' n' \ \ (p'::'a::len word) + mask n' \ p + mask n" apply (cases "LENGTH('a) \ n") apply (frule (1) is_aligned_over_length) apply (drule mask_over_length) apply clarsimp apply (simp add: not_le) apply (simp add: word_le_nat_alt unat_plus_simple) apply (subst unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask)+ apply (subst (asm) unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask) apply (clarsimp simp: is_aligned_def dvd_def) apply (rename_tac k k') apply (thin_tac "unat p = x" for p x)+ apply (subst Suc_le_mono[symmetric]) apply (simp only: Suc_2p_unat_mask) apply (drule le_imp_less_Suc, subst (asm) Suc_2p_unat_mask, assumption) apply (erule (1) power_2_mult_step_le) done lemma mask_mono: "sz' \ sz \ mask sz' \ (mask sz :: 'a::len word)" by (simp add: le_mask_iff shiftr_mask_le) lemma aligned_mask_disjoint: "\ is_aligned (a :: 'a :: len word) n; b \ mask n \ \ a && b = 0" by word_eqI_solve lemma word_and_or_mask_aligned: "\ is_aligned a n; b \ mask n \ \ a + b = a || b" by (simp add: aligned_mask_disjoint word_plus_and_or_coroll) lemmas word_and_or_mask_aligned2 = word_and_or_mask_aligned[where a=b and b=a for a b, simplified add.commute word_bool_alg.disj.commute] lemma is_aligned_ucastI: "is_aligned w n \ is_aligned (ucast w) n" by (clarsimp simp: word_eqI_simps) lemma ucast_le_maskI: "a \ mask n \ UCAST('a::len \ 'b::len) a \ mask n" by (metis and_mask_eq_iff_le_mask ucast_and_mask) lemma ucast_add_mask_aligned: "\ a \ mask n; is_aligned b n \ \ UCAST ('a::len \ 'b::len) (a + b) = ucast a + ucast b" by (metis is_aligned_ucastI ucast_le_maskI ucast_or_distrib word_and_or_mask_aligned2) lemma ucast_shiftl: "LENGTH('b) \ LENGTH ('a) \ UCAST ('a::len \ 'b::len) x << n = ucast (x << n)" by word_eqI_solve lemma ucast_leq_mask: - "LENGTH('a) \ n \ ucast (x::'a::len0 word) \ mask n" + "LENGTH('a) \ n \ ucast (x::'a::len word) \ mask n" by (clarsimp simp: le_mask_high_bits word_size nth_ucast) lemma shiftl_inj: "\ x << n = y << n; x \ mask (LENGTH('a)-n); y \ mask (LENGTH('a)-n) \ \ x = (y :: 'a :: len word)" apply word_eqI apply (rename_tac n') apply (case_tac "LENGTH('a) - n \ n'", simp) by (metis add.commute add.right_neutral diff_add_inverse le_diff_conv linorder_not_less zero_order(1)) lemma distinct_word_add_ucast_shift_inj: "\ p + (UCAST('a::len \ 'b::len) off << n) = p' + (ucast off' << n); is_aligned p n'; is_aligned p' n'; n' = n + LENGTH('a); n' < LENGTH('b) \ \ p' = p \ off' = off" apply (simp add: word_and_or_mask_aligned le_mask_shiftl_le_mask[where n="LENGTH('a)"] ucast_leq_mask) apply (simp add: is_aligned_nth) apply (rule conjI; word_eqI) apply (metis add.commute test_bit_conj_lt diff_add_inverse le_diff_conv nat_less_le) apply (rename_tac i) apply (erule_tac x="i+n" in allE) apply simp done lemma aligned_add_mask_lessD: "\ x + mask n < y; is_aligned x n \ \ x < y" for y::"'a::len word" by (metis is_aligned_no_overflow' mask_2pm1 order_le_less_trans) lemma aligned_add_mask_less_eq: "\ is_aligned x n; is_aligned y n; n < LENGTH('a) \ \ (x + mask n < y) = (x < y)" for y::"'a::len word" using aligned_add_mask_lessD is_aligned_add_step_le word_le_not_less by blast lemma word_upto_Nil: "y < x \ [x .e. y ::'a::len word] = []" by (simp add: upto_enum_red not_le word_less_nat_alt) lemma word_enum_decomp_elem: assumes "[x .e. (y ::'a::len word)] = as @ a # bs" shows "x \ a \ a \ y" proof - have "set as \ set [x .e. y] \ a \ set [x .e. y]" using assms by (auto dest: arg_cong[where f=set]) then show ?thesis by auto qed lemma max_word_not_less[simp]: "\ max_word < x" by (simp add: not_less) lemma word_enum_prefix: "[x .e. (y ::'a::len word)] = as @ a # bs \ as = (if x < a then [x .e. a - 1] else [])" apply (induct as arbitrary: x; clarsimp) apply (case_tac "x < y") prefer 2 apply (case_tac "x = y", simp) apply (simp add: not_less) apply (drule (1) dual_order.not_eq_order_implies_strict) apply (simp add: word_upto_Nil) apply (simp add: word_upto_Cons_eq) apply (case_tac "x < y") prefer 2 apply (case_tac "x = y", simp) apply (simp add: not_less) apply (drule (1) dual_order.not_eq_order_implies_strict) apply (simp add: word_upto_Nil) apply (clarsimp simp: word_upto_Cons_eq) apply (frule word_enum_decomp_elem) apply clarsimp apply (rule conjI) prefer 2 apply (subst word_Suc_le[symmetric]; clarsimp) apply (drule meta_spec) apply (drule (1) meta_mp) apply clarsimp apply (rule conjI; clarsimp) apply (subst (2) word_upto_Cons_eq) apply unat_arith apply simp done lemma word_enum_decomp_set: "[x .e. (y ::'a::len word)] = as @ a # bs \ a \ set as" by (metis distinct_append distinct_enum_upto' not_distinct_conv_prefix) lemma word_enum_decomp: assumes "[x .e. (y ::'a::len word)] = as @ a # bs" shows "x \ a \ a \ y \ a \ set as \ (\z \ set as. x \ z \ z \ y)" proof - from assms have "set as \ set [x .e. y] \ a \ set [x .e. y]" by (auto dest: arg_cong[where f=set]) with word_enum_decomp_set[OF assms] show ?thesis by auto qed lemma of_nat_unat_le_mask_ucast: "\of_nat (unat t) = w; t \ mask LENGTH('a)\ \ t = UCAST('a::len \ 'b::len) w" by (clarsimp simp: ucast_nat_def ucast_ucast_mask simp flip: and_mask_eq_iff_le_mask) lemma fold_eq_0_to_bool: "(v = 0) = (\ to_bool v)" by (simp add: to_bool_def) lemma less_diff_gt0: "a < b \ (0 :: 'a :: len word) < b - a" by unat_arith lemma unat_plus_gt: "unat ((a :: 'a :: len word) + b) \ unat a + unat b" by (clarsimp simp: unat_plus_if_size) lemma const_less: "\ (a :: 'a :: len word) - 1 < b; a \ b \ \ a < b" by (metis less_1_simp word_le_less_eq) lemma add_mult_aligned_neg_mask: "m && (2 ^ n - 1) = 0 \ (x + y * m) && ~~(mask n) = (x && ~~(mask n)) + y * m" by (metis Groups.add_ac(2) is_aligned_mask mask_def mask_eqs(5) mask_out_add_aligned mult_zero_right shiftl_1 word_bw_comms(1) word_log_esimps(1)) lemma unat_of_nat_minus_1: "\ n < 2 ^ LENGTH('a); n \ 0 \ \ unat ((of_nat n:: 'a :: len word) - 1) = n - 1" by (simp add: unat_eq_of_nat) lemma word_eq_zeroI: "a \ a - 1 \ a = 0" for a :: "'a :: len word" by (simp add: word_must_wrap) lemma word_add_format: "(-1 :: 'a :: len word) + b + c = b + (c - 1)" by simp lemma upto_enum_word_nth: "\ i \ j; k \ unat (j - i) \ \ [i .e. j] ! k = i + of_nat k" apply (clarsimp simp: upto_enum_def nth_append) apply (clarsimp simp: word_le_nat_alt[symmetric]) apply (rule conjI, clarsimp) apply (subst toEnum_of_nat, unat_arith) apply unat_arith apply (clarsimp simp: not_less unat_sub[symmetric]) apply unat_arith done lemma upto_enum_step_nth: "\ a \ c; n \ unat ((c - a) div (b - a)) \ \ [a, b .e. c] ! n = a + of_nat n * (b - a)" by (clarsimp simp: upto_enum_step_def not_less[symmetric] upto_enum_word_nth) lemma upto_enum_inc_1_len: "a < - 1 \ [(0 :: 'a :: len word) .e. 1 + a] = [0 .e. a] @ [1 + a]" apply (simp add: upto_enum_word) apply (subgoal_tac "unat (1+a) = 1 + unat a") apply simp apply (subst unat_plus_simple[THEN iffD1]) apply (metis add.commute no_plus_overflow_neg olen_add_eqv) apply unat_arith done lemma neg_mask_add: "y && mask n = 0 \ x + y && ~~(mask n) = (x && ~~(mask n)) + y" by (clarsimp simp: mask_out_sub_mask mask_eqs(7)[symmetric] mask_twice) lemma shiftr_shiftl_shiftr[simp]: "(x :: 'a :: len word) >> a << a >> a = x >> a" by word_eqI_solve lemma add_right_shift: "\ x && mask n = 0; y && mask n = 0; x \ x + y \ \ (x + y :: ('a :: len) word) >> n = (x >> n) + (y >> n)" apply (simp add: no_olen_add_nat is_aligned_mask[symmetric]) apply (simp add: unat_arith_simps shiftr_div_2n' split del: if_split) apply (subst if_P) apply (erule order_le_less_trans[rotated]) apply (simp add: add_mono) apply (simp add: shiftr_div_2n' is_aligned_def) done lemma sub_right_shift: "\ x && mask n = 0; y && mask n = 0; y \ x \ \ (x - y) >> n = (x >> n :: 'a :: len word) - (y >> n)" using add_right_shift[where x="x - y" and y=y and n=n] by (simp add: aligned_sub_aligned is_aligned_mask[symmetric] word_sub_le) lemma and_and_mask_simple: "y && mask n = mask n \ (x && y) && mask n = x && mask n" by (simp add: word_bool_alg.conj.assoc) lemma and_and_mask_simple_not: "y && mask n = 0 \ (x && y) && mask n = 0" by (simp add: word_bool_alg.conj.assoc) lemma word_and_le': "b \ c \ (a :: 'a :: len word) && b \ c" by (metis word_and_le1 order_trans) lemma word_and_less': "b < c \ (a :: 'a :: len word) && b < c" by (metis word_and_le1 xtr7) lemma shiftr_w2p: "x < LENGTH('a) \ 2 ^ x = (2 ^ (LENGTH('a) - 1) >> (LENGTH('a) - 1 - x) :: 'a :: len word)" by word_eqI_solve lemma t2p_shiftr: "\ b \ a; a < LENGTH('a) \ \ (2 :: 'a :: len word) ^ a >> b = 2 ^ (a - b)" by word_eqI_solve lemma scast_1[simp]: "scast (1 :: 'a :: len signed word) = (1 :: 'a word)" by simp lemma ucast_ucast_mask_eq: "\ UCAST('a::len \ 'b::len) x = y; x && mask LENGTH('b) = x \ \ x = ucast y" by word_eqI_solve lemma ucast_up_eq: "\ ucast x = (ucast y::'b::len word); LENGTH('a) \ LENGTH ('b) \ \ ucast x = (ucast y::'a::len word)" by word_eqI_solve lemma ucast_up_neq: "\ ucast x \ (ucast y::'b::len word); LENGTH('b) \ LENGTH ('a) \ \ ucast x \ (ucast y::'a::len word)" by (fastforce dest: ucast_up_eq) lemma mask_AND_less_0: "\ x && mask n = 0; m \ n \ \ x && mask m = 0" by (metis mask_twice2 word_and_notzeroD) lemma mask_len_id [simp]: "(x :: 'a :: len word) && mask LENGTH('a) = x" using uint_lt2p [of x] by (simp add: mask_eq_iff) lemma scast_ucast_down_same: "LENGTH('b) \ LENGTH('a) \ SCAST('a \ 'b) = UCAST('a::len \ 'b::len)" by (simp add: down_cast_same is_down) lemma word_aligned_0_sum: "\ a + b = 0; is_aligned (a :: 'a :: len word) n; b \ mask n; n < LENGTH('a) \ \ a = 0 \ b = 0" by (simp add: word_plus_and_or_coroll aligned_mask_disjoint word_or_zero) lemma mask_eq1_nochoice: "\ LENGTH('a) > 1; (x :: 'a :: len word) && 1 = x \ \ x = 0 \ x = 1" by (metis word_and_1) lemma pow_mono_leq_imp_lt: "x \ y \ x < 2 ^ y" by (simp add: le_less_trans) lemma unat_of_nat_ctz_mw: "unat (of_nat (word_ctz (w :: 'a :: len word)) :: 'a :: len word) = word_ctz w" using word_ctz_le[where w=w, simplified] unat_of_nat_eq[where x="word_ctz w" and 'a="'a"] pow_mono_leq_imp_lt by simp lemma unat_of_nat_ctz_smw: "unat (of_nat (word_ctz (w :: 'a :: len word)) :: 'a :: len sword) = word_ctz w" using word_ctz_le[where w=w, simplified] unat_of_nat_eq[where x="word_ctz w" and 'a="'a"] pow_mono_leq_imp_lt by (metis le_unat_uoi le_unat_uoi linorder_neqE_nat nat_less_le scast_of_nat word_unat.Rep_inverse) lemma shiftr_and_eq_shiftl: "(w >> n) && x = y \ w && (x << n) = (y << n)" for y :: "'a:: len word" by (smt and_not_mask is_aligned_neg_mask_eq is_aligned_shift shift_over_ao_dists(4) word_bool_alg.conj_assoc word_bw_comms(1)) lemma neg_mask_combine: "~~(mask a) && ~~(mask b) = ~~(mask (max a b))" by (auto simp: word_ops_nth_size word_size intro!: word_eqI) lemma neg_mask_twice: "x && ~~(mask n) && ~~(mask m) = x && ~~(mask (max n m))" by (metis neg_mask_combine) lemma multiple_mask_trivia: "n \ m \ (x && ~~(mask n)) + (x && mask n && ~~(mask m)) = x && ~~(mask m)" apply (rule trans[rotated], rule_tac w="mask n" in word_plus_and_or_coroll2) apply (simp add: word_bw_assocs word_bw_comms word_bw_lcs neg_mask_twice max_absorb2) done lemma add_mask_lower_bits': "\ len = LENGTH('a); is_aligned (x :: 'a :: len word) n; \n' \ n. n' < len \ \ p !! n' \ \ x + p && ~~(mask n) = x" using add_mask_lower_bits by auto lemma neg_mask_in_mask_range: "is_aligned ptr bits \ (ptr' && ~~(mask bits) = ptr) = (ptr' \ mask_range ptr bits)" apply (erule is_aligned_get_word_bits) apply (rule iffI) apply (drule sym) apply (simp add: word_and_le2) apply (subst word_plus_and_or_coroll, word_eqI_solve) apply (metis le_word_or2 neg_mask_add_mask word_bool_alg.conj.right_idem) apply clarsimp apply (smt add.right_neutral eq_iff is_aligned_neg_mask_eq mask_out_add_aligned neg_mask_mono_le word_and_not) apply (simp add: power_overflow mask_def) done lemma aligned_offset_in_range: "\ is_aligned (x :: 'a :: len word) m; y < 2 ^ m; is_aligned p n; n \ m; n < LENGTH('a) \ \ (x + y \ {p .. p + mask n}) = (x \ mask_range p n)" apply (simp only: is_aligned_add_or flip: neg_mask_in_mask_range) by (metis less_mask_eq mask_subsume) lemma mask_range_to_bl': "\ is_aligned (ptr :: 'a :: len word) bits; bits < LENGTH('a) \ \ mask_range ptr bits = {x. take (LENGTH('a) - bits) (to_bl x) = take (LENGTH('a) - bits) (to_bl ptr)}" apply (rule set_eqI, rule iffI) apply clarsimp apply (subgoal_tac "\y. x = ptr + y \ y < 2 ^ bits") apply clarsimp apply (subst is_aligned_add_conv) apply assumption apply simp apply simp apply (rule_tac x="x - ptr" in exI) apply (simp add: add_diff_eq[symmetric]) apply (simp only: word_less_sub_le[symmetric]) apply (rule word_diff_ls') apply (simp add: field_simps mask_def) apply assumption apply simp apply (subgoal_tac "\y. y < 2 ^ bits \ to_bl (ptr + y) = to_bl x") apply clarsimp apply (rule conjI) apply (erule(1) is_aligned_no_wrap') apply (simp only: add_diff_eq[symmetric] mask_def) apply (rule word_plus_mono_right) apply simp apply (erule is_aligned_no_wrap') apply simp apply (rule_tac x="of_bl (drop (LENGTH('a) - bits) (to_bl x))" in exI) apply (rule context_conjI) apply (rule order_less_le_trans [OF of_bl_length]) apply simp apply simp apply (subst is_aligned_add_conv) apply assumption apply simp apply (drule sym) apply (simp add: word_rep_drop) done lemma mask_range_to_bl: "is_aligned (ptr :: 'a :: len word) bits \ mask_range ptr bits = {x. take (LENGTH('a) - bits) (to_bl x) = take (LENGTH('a) - bits) (to_bl ptr)}" apply (erule is_aligned_get_word_bits) apply (erule(1) mask_range_to_bl') apply (rule set_eqI) apply (simp add: power_overflow mask_def) done lemma aligned_mask_range_cases: "\ is_aligned (p :: 'a :: len word) n; is_aligned (p' :: 'a :: len word) n' \ \ mask_range p n \ mask_range p' n' = {} \ mask_range p n \ mask_range p' n' \ mask_range p n \ mask_range p' n'" apply (simp add: mask_range_to_bl) apply (rule Meson.disj_comm, rule disjCI) apply (erule nonemptyE) apply simp apply (subgoal_tac "(\n''. LENGTH('a) - n = (LENGTH('a) - n') + n'') \ (\n''. LENGTH('a) - n' = (LENGTH('a) - n) + n'')") apply (fastforce simp: take_add) apply arith done lemma aligned_mask_range_offset_subset: assumes al: "is_aligned (ptr :: 'a :: len word) sz" and al': "is_aligned x sz'" and szv: "sz' \ sz" and xsz: "x < 2 ^ sz" shows "mask_range (ptr+x) sz' \ mask_range ptr sz" using al proof (rule is_aligned_get_word_bits) assume p0: "ptr = 0" and szv': "LENGTH ('a) \ sz" then have "(2 ::'a word) ^ sz = 0" by simp show ?thesis using p0 by (simp add: \2 ^ sz = 0\ mask_def) next assume szv': "sz < LENGTH('a)" hence blah: "2 ^ (sz - sz') < (2 :: nat) ^ LENGTH('a)" using szv by auto show ?thesis using szv szv' apply (intro range_subsetI) apply (rule is_aligned_no_wrap' [OF al xsz]) apply (simp only: flip: add_diff_eq add_mask_fold) apply (subst add.assoc, rule word_plus_mono_right) using al' is_aligned_add_less_t2n xsz apply fastforce apply (simp add: field_simps szv al is_aligned_no_overflow) done qed lemma aligned_mask_diff: "\ is_aligned (dest :: 'a :: len word) bits; is_aligned (ptr :: 'a :: len word) sz; bits \ sz; sz < LENGTH('a); dest < ptr \ \ mask bits + dest < ptr" apply (frule_tac p' = ptr in aligned_mask_range_cases, assumption) apply (elim disjE) apply (drule_tac is_aligned_no_overflow_mask, simp)+ apply (simp add: algebra_split_simps word_le_not_less) apply (drule is_aligned_no_overflow_mask; fastforce) by (simp add: aligned_add_mask_less_eq is_aligned_weaken algebra_split_simps) lemma aligned_mask_ranges_disjoint: "\ is_aligned (p :: 'a :: len word) n; is_aligned (p' :: 'a :: len word) n'; p && ~~(mask n') \ p'; p' && ~~(mask n) \ p \ \ mask_range p n \ mask_range p' n' = {}" using aligned_mask_range_cases by (auto simp: neg_mask_in_mask_range) lemma aligned_mask_ranges_disjoint2: "\ is_aligned p n; is_aligned ptr bits; n \ m; n < size p; m \ bits; (\y < 2 ^ (n - m). p + (y << m) \ mask_range ptr bits) \ \ mask_range p n \ mask_range ptr bits = {}" apply safe apply (simp only: flip: neg_mask_in_mask_range) apply (drule_tac x="x && mask n >> m" in spec) apply (clarsimp simp: shiftr_less_t2n and_mask_less_size wsst_TYs multiple_mask_trivia word_bw_assocs neg_mask_twice max_absorb2 shiftr_shiftl1) done lemma leq_mask_shift: "(x :: 'a :: len word) \ mask (low_bits + high_bits) \ (x >> low_bits) \ mask high_bits" by (simp add: le_mask_iff shiftr_shiftr) lemma ucast_ucast_eq_mask_shift: "(x :: 'a :: len word) \ mask (low_bits + LENGTH('b)) \ ucast((ucast (x >> low_bits)) :: 'b :: len word) = x >> low_bits" by (meson and_mask_eq_iff_le_mask eq_ucast_ucast_eq not_le_imp_less shiftr_less_t2n' ucast_ucast_len) lemma const_le_unat: "\ b < 2 ^ LENGTH('a); of_nat b \ a \ \ b \ unat (a :: 'a :: len word)" by (clarsimp simp: word_le_def uint_nat of_nat_inverse) lemma upt_enum_offset_trivial: "\ x < 2 ^ LENGTH('a) - 1 ; n \ unat x \ \ ([(0 :: 'a :: len word) .e. x] ! n) = of_nat n" apply (induct x arbitrary: n) apply simp by (simp add: upto_enum_word_nth) lemma word_le_mask_out_plus_2sz: "x \ (x && ~~(mask sz)) + 2 ^ sz - 1" by (metis add_diff_eq word_neg_and_le) lemma ucast_add: "ucast (a + (b :: 'a :: len word)) = ucast a + (ucast b :: ('a signed word))" apply (case_tac "LENGTH('a) = 1") apply (clarsimp simp: ucast_def) apply (metis (hide_lams, mono_tags) One_nat_def len_signed plus_word.abs_eq uint_word_arith_bintrs(1) word_ubin.Abs_norm) apply (clarsimp simp: ucast_def) apply (metis le_refl len_signed plus_word.abs_eq uint_word_arith_bintrs(1) wi_bintr) done lemma ucast_minus: "ucast (a - (b :: 'a :: len word)) = ucast a - (ucast b :: ('a signed word))" apply (insert ucast_add[where a=a and b="-b"]) apply (metis (no_types, hide_lams) add_diff_eq diff_add_cancel ucast_add) done lemma scast_ucast_add_one [simp]: "scast (ucast (x :: 'a::len word) + (1 :: 'a signed word)) = x + 1" apply (subst ucast_1[symmetric]) apply (subst ucast_add[symmetric]) apply clarsimp done lemma word_and_le_plus_one: "a > 0 \ (x :: 'a :: len word) && (a - 1) < a" by (simp add: gt0_iff_gem1 word_and_less') lemma unat_of_ucast_then_shift_eq_unat_of_shift[simp]: "LENGTH('b) \ LENGTH('a) \ unat ((ucast (x :: 'a :: len word) :: 'b :: len word) >> n) = unat (x >> n)" by (simp add: shiftr_div_2n' unat_ucast_up_simp) lemma unat_of_ucast_then_mask_eq_unat_of_mask[simp]: "LENGTH('b) \ LENGTH('a) \ unat ((ucast (x :: 'a :: len word) :: 'b :: len word) && mask m) = unat (x && mask m)" by (metis ucast_and_mask unat_ucast_up_simp) lemma small_powers_of_2: "x \ 3 \ x < 2 ^ (x - 1)" by (induct x; simp add: suc_le_pow_2) lemma word_clz_sint_upper[simp]: "LENGTH('a) \ 3 \ sint (of_nat (word_clz (w :: 'a :: len word)) :: 'a sword) \ int (LENGTH('a))" using small_powers_of_2 by (smt One_nat_def diff_less le_less_trans len_gt_0 len_signed lessI n_less_equal_power_2 not_msb_from_less of_nat_mono sint_eq_uint uint_nat unat_of_nat_eq unat_power_lower word_clz_max word_of_nat_less wsst_TYs(3)) lemma word_clz_sint_lower[simp]: "LENGTH('a) \ 3 \ - sint (of_nat (word_clz (w :: 'a :: len word)) :: 'a signed word) \ int (LENGTH('a))" apply (subst sint_eq_uint) using small_powers_of_2 uint_nat apply (simp add: order_le_less_trans[OF word_clz_max] not_msb_from_less word_of_nat_less word_size) by (simp add: uint_nat) lemma shiftr_less_t2n3: "\ (2 :: 'a word) ^ (n + m) = 0; m < LENGTH('a) \ \ (x :: 'a :: len word) >> n < 2 ^ m" by (fastforce intro: shiftr_less_t2n' simp: mask_def power_overflow) lemma unat_shiftr_le_bound: "\ 2 ^ (LENGTH('a :: len) - n) - 1 \ bnd; 0 < n \ \ unat ((x :: 'a word) >> n) \ bnd" using less_not_refl3 le_step_down_nat le_trans less_or_eq_imp_le word_shiftr_lt by (metis (no_types, lifting)) lemma shiftr_eqD: "\ x >> n = y >> n; is_aligned x n; is_aligned y n \ \ x = y" by (metis is_aligned_shiftr_shiftl) lemma word_shiftr_shiftl_shiftr_eq_shiftr: "a \ b \ (x :: 'a :: len word) >> a << b >> b = x >> a" by (simp add: mask_shift multi_shift_simps(5) shiftr_shiftr) lemma of_int_uint_ucast: "of_int (uint (x :: 'a::len word)) = (ucast x :: 'b::len word)" by (simp add: ucast_def word_of_int) lemma mod_mask_drop: "\ m = 2 ^ n; 0 < m; mask n && msk = mask n \ \ (x mod m) && msk = x mod m" by (simp add: word_mod_2p_is_mask word_bw_assocs) lemma mask_eq_ucast_eq: "\ x && mask LENGTH('a) = (x :: ('c :: len word)); LENGTH('a) \ LENGTH('b)\ \ ucast (ucast x :: ('a :: len word)) = (ucast x :: ('b :: len word))" by (metis ucast_and_mask ucast_id ucast_ucast_mask ucast_up_eq) lemma of_nat_less_t2n: "of_nat i < (2 :: ('a :: len) word) ^ n \ n < LENGTH('a) \ unat (of_nat i :: 'a word) < 2 ^ n" by (metis order_less_trans p2_gt_0 unat_less_power word_neq_0_conv) lemma two_power_increasing_less_1: "\ n \ m; m \ LENGTH('a) \ \ (2 :: 'a :: len word) ^ n - 1 \ 2 ^ m - 1" by (metis diff_diff_cancel le_m1_iff_lt less_imp_diff_less p2_gt_0 two_power_increasing word_1_le_power word_le_minus_mono_left word_less_sub_1) lemma word_sub_mono4: "\ y + x \ z + x; y \ y + x; z \ z + x \ \ y \ z" for y :: "'a :: len word" by (simp add: word_add_le_iff2) lemma eq_or_less_helperD: "\ n = unat (2 ^ m - 1 :: 'a :: len word) \ n < unat (2 ^ m - 1 :: 'a word); m < LENGTH('a) \ \ n < 2 ^ m" by (meson le_less_trans nat_less_le unat_less_power word_power_less_1) lemma mask_sub: "n \ m \ mask m - mask n = mask m && ~~(mask n)" by (metis (no_types) AND_NOT_mask_plus_AND_mask_eq add_right_imp_eq diff_add_cancel and_mask_eq_iff_shiftr_0 shiftr_mask_le word_bool_alg.conj.commute) lemma neg_mask_diff_bound: "sz'\ sz \ (ptr && ~~(mask sz')) - (ptr && ~~(mask sz)) \ 2 ^ sz - 2 ^ sz'" (is "_ \ ?lhs \ ?rhs") proof - assume lt: "sz' \ sz" hence "?lhs = ptr && (mask sz && ~~(mask sz'))" by (metis add_diff_cancel_left' multiple_mask_trivia) also have "\ \ ?rhs" using lt by (metis (mono_tags) add_diff_eq diff_eq_eq eq_iff mask_2pm1 mask_sub word_and_le') finally show ?thesis by simp qed lemma mask_range_subsetD: "\ p' \ mask_range p n; x' \ mask_range p' n'; n' \ n; is_aligned p n; is_aligned p' n' \ \ x' \ mask_range p n" using aligned_mask_step by fastforce lemma add_mult_in_mask_range: "\ is_aligned (base :: 'a :: len word) n; n < LENGTH('a); bits \ n; x < 2 ^ (n - bits) \ \ base + x * 2^bits \ mask_range base n" by (simp add: is_aligned_no_wrap' mask_2pm1 nasty_split_lt word_less_power_trans2 word_plus_mono_right) lemma of_bl_length2: "length xs + c < LENGTH('a) \ of_bl xs * 2^c < (2::'a::len word) ^ (length xs + c)" by (simp add: of_bl_length word_less_power_trans2) lemma mask_out_eq_0: "\ idx < 2 ^ sz; sz < LENGTH('a) \ \ (of_nat idx :: 'a :: len word) && ~~(mask sz) = 0" by (simp add: Word_Lemmas.of_nat_power less_mask_eq mask_eq_0_eq_x) lemma is_aligned_neg_mask_eq': "is_aligned ptr sz = (ptr && ~~(mask sz) = ptr)" using is_aligned_mask mask_eq_0_eq_x by blast lemma neg_mask_mask_unat: "sz < LENGTH('a) \ unat ((ptr :: 'a :: len word) && ~~(mask sz)) + unat (ptr && mask sz) = unat ptr" by (metis AND_NOT_mask_plus_AND_mask_eq unat_plus_simple word_and_le2) lemma unat_pow_le_intro: "LENGTH('a) \ n \ unat (x :: 'a :: len word) < 2 ^ n" by (metis lt2p_lem not_le of_nat_le_iff of_nat_numeral semiring_1_class.of_nat_power uint_nat) lemma unat_shiftl_less_t2n: "\ unat (x :: 'a :: len word) < 2 ^ (m - n); m < LENGTH('a) \ \ unat (x << n) < 2 ^ m" by (metis (no_types) Word_Lemmas.of_nat_power diff_le_self le_less_trans shiftl_less_t2n unat_less_power word_unat.Rep_inverse) lemma unat_is_aligned_add: "\ is_aligned p n; unat d < 2 ^ n \ \ unat (p + d && mask n) = unat d \ unat (p + d && ~~(mask n)) = unat p" by (metis add.right_neutral and_mask_eq_iff_le_mask and_not_mask le_mask_iff mask_add_aligned mask_out_add_aligned mult_zero_right shiftl_t2n shiftr_le_0) lemma unat_shiftr_shiftl_mask_zero: "\ c + a \ LENGTH('a) + b ; c < LENGTH('a) \ \ unat (((q :: 'a :: len word) >> a << b) && ~~(mask c)) = 0" by (fastforce intro: unat_is_aligned_add[where p=0 and n=c, simplified, THEN conjunct2] unat_shiftl_less_t2n unat_shiftr_less_t2n unat_pow_le_intro) lemmas of_nat_ucast = ucast_of_nat[symmetric] lemma shift_then_mask_eq_shift_low_bits: "x \ mask (low_bits + high_bits) \ (x >> low_bits) && mask high_bits = x >> low_bits" by (simp add: leq_mask_shift le_mask_imp_and_mask) lemma leq_low_bits_iff_zero: "\ x \ mask (low bits + high bits); x >> low_bits = 0 \ \ (x && mask low_bits = 0) = (x = 0)" using and_mask_eq_iff_shiftr_0 by force lemma unat_less_iff: "\ unat (a :: 'a :: len word) = b; c < 2 ^ LENGTH('a) \ \ (a < of_nat c) = (b < c)" using unat_ucast_less_no_overflow_simp by blast lemma is_aligned_no_overflow3: "\ is_aligned (a :: 'a :: len word) n; n < LENGTH('a); b < 2 ^ n; c \ 2 ^ n; b < c \ \ a + b \ a + (c - 1)" by (meson is_aligned_no_wrap' le_m1_iff_lt not_le word_less_sub_1 word_plus_mono_right) lemma mask_add_aligned_right: "is_aligned p n \ (q + p) && mask n = q && mask n" by (simp add: mask_add_aligned add.commute) lemma leq_high_bits_shiftr_low_bits_leq_bits_mask: "x \ mask high_bits \ (x :: 'a :: len word) << low_bits \ mask (low_bits + high_bits)" by (metis le_mask_shiftl_le_mask) lemma from_to_bool_last_bit: "from_bool (to_bool (x && 1)) = x && 1" by (metis from_bool_to_bool_iff word_and_1) lemma word_two_power_neg_ineq: "2 ^ m \ (0 :: 'a word) \ 2 ^ n \ - (2 ^ m :: 'a :: len word)" apply (cases "n < LENGTH('a)"; simp add: power_overflow) apply (cases "m < LENGTH('a)"; simp add: power_overflow) apply (simp add: word_le_nat_alt unat_minus word_size) apply (cases "LENGTH('a)"; simp) apply (simp add: less_Suc_eq_le) apply (drule power_increasing[where a=2 and n=n] power_increasing[where a=2 and n=m], simp)+ apply (drule(1) add_le_mono) apply simp done lemma unat_shiftl_absorb: "\ x \ 2 ^ p; p + k < LENGTH('a) \ \ unat (x :: 'a :: len word) * 2 ^ k = unat (x * 2 ^ k)" by (smt add_diff_cancel_right' add_lessD1 le_add2 le_less_trans mult.commute nat_le_power_trans unat_lt2p unat_mult_lem unat_power_lower word_le_nat_alt) lemma word_plus_mono_right_split: "\ unat ((x :: 'a :: len word) && mask sz) + unat z < 2 ^ sz; sz < LENGTH('a) \ \ x \ x + z" apply (subgoal_tac "(x && ~~(mask sz)) + (x && mask sz) \ (x && ~~(mask sz)) + ((x && mask sz) + z)") apply (simp add:word_plus_and_or_coroll2 field_simps) apply (rule word_plus_mono_right) apply (simp add: less_le_trans no_olen_add_nat) using Word_Lemmas.of_nat_power is_aligned_no_wrap' by force lemma mul_not_mask_eq_neg_shiftl: "~~(mask n) = -1 << n" by (simp add: NOT_mask shiftl_t2n) lemma shiftr_mul_not_mask_eq_and_not_mask: "(x >> n) * ~~(mask n) = - (x && ~~(mask n))" by (metis NOT_mask and_not_mask mult_minus_left semiring_normalization_rules(7) shiftl_t2n) lemma mask_eq_n1_shiftr: "n \ LENGTH('a) \ (mask n :: 'a :: len word) = -1 >> (LENGTH('a) - n)" by (metis diff_diff_cancel eq_refl mask_full shiftr_mask2) lemma is_aligned_mask_out_add_eq: "is_aligned p n \ (p + x) && ~~(mask n) = p + (x && ~~(mask n))" by (simp add: mask_out_sub_mask mask_add_aligned) lemmas is_aligned_mask_out_add_eq_sub = is_aligned_mask_out_add_eq[where x="a - b" for a b, simplified field_simps] lemma aligned_bump_down: "is_aligned x n \ (x - 1) && ~~(mask n) = x - 2 ^ n" by (drule is_aligned_mask_out_add_eq[where x="-1"]) (simp add: NOT_mask) lemma unat_2tp_if: "unat (2 ^ n :: ('a :: len) word) = (if n < LENGTH ('a) then 2 ^ n else 0)" by (split if_split, simp_all add: power_overflow) lemma mask_of_mask: "mask (n::nat) && mask (m::nat) = mask (min m n)" by word_eqI_solve lemma unat_signed_ucast_less_ucast: "LENGTH('a) \ LENGTH('b) \ unat (ucast (x :: 'a :: len word) :: 'b :: len signed word) = unat x" by (simp add: unat_ucast_up_simp) lemma toEnum_of_ucast: "LENGTH('b) \ LENGTH('a) \ (toEnum (unat (b::'b :: len word))::'a :: len word) = of_nat (unat b)" by (simp add: unat_pow_le_intro) lemmas unat_ucast_mask = unat_ucast_eq_unat_and_mask[where w=a for a] lemma t2n_mask_eq_if: "2 ^ n && mask m = (if n < m then 2 ^ n else 0)" by (rule word_eqI, auto simp add: word_size nth_w2p split: if_split) lemma unat_ucast_le: "unat (ucast (x :: 'a :: len word) :: 'b :: len word) \ unat x" by (simp add: ucast_nat_def word_unat_less_le) lemma ucast_le_up_down_iff: "\ LENGTH('a) \ LENGTH('b); (x :: 'b :: len word) \ ucast (max_word :: 'a :: len word) \ \ (ucast x \ (y :: 'a word)) = (x \ ucast y)" using le_max_word_ucast_id ucast_le_ucast by metis lemma ucast_ucast_mask_shift: "a \ LENGTH('a) + b \ ucast (ucast (p && mask a >> b) :: 'a :: len word) = p && mask a >> b" by (metis add.commute le_mask_iff shiftr_mask_le ucast_ucast_eq_mask_shift word_and_le') lemma unat_ucast_mask_shift: "a \ LENGTH('a) + b \ unat (ucast (p && mask a >> b) :: 'a :: len word) = unat (p && mask a >> b)" by (metis linear ucast_ucast_mask_shift unat_ucast_up_simp) lemma mask_overlap_zero: "a \ b \ (p && mask a) && ~~(mask b) = 0" by (metis NOT_mask_AND_mask mask_lower_twice2 max_def) lemma mask_shifl_overlap_zero: "a + c \ b \ (p && mask a << c) && ~~(mask b) = 0" by (metis and_not_mask le_mask_iff mask_shiftl_decompose shiftl_0 shiftl_over_and_dist shiftr_mask_le word_and_le' word_bool_alg.conj_commute) lemma mask_overlap_zero': "a \ b \ (p && ~~(mask a)) && mask b = 0" using mask_AND_NOT_mask mask_AND_less_0 by blast lemma mask_rshift_mult_eq_rshift_lshift: "((a :: 'a :: len word) >> b) * (1 << c) = (a >> b << c)" by (simp add: shiftl_t2n) lemma shift_alignment: "a \ b \ is_aligned (p >> a << a) b" using is_aligned_shift is_aligned_weaken by blast lemma mask_split_sum_twice: "a \ b \ (p && ~~(mask a)) + ((p && mask a) && ~~(mask b)) + (p && mask b) = p" by (simp add: add.commute multiple_mask_trivia word_bool_alg.conj_commute word_bool_alg.conj_left_commute word_plus_and_or_coroll2) lemma mask_shift_eq_mask_mask: "(p && mask a >> b << b) = (p && mask a) && ~~(mask b)" by (simp add: and_not_mask) lemma mask_shift_sum: "\ a \ b; unat n = unat (p && mask b) \ \ (p && ~~(mask a)) + (p && mask a >> b) * (1 << b) + n = (p :: 'a :: len word)" by (metis and_not_mask mask_rshift_mult_eq_rshift_lshift mask_split_sum_twice word_unat.Rep_eqD) lemma is_up_compose: "\ is_up uc; is_up uc' \ \ is_up (uc' \ uc)" unfolding is_up_def by (simp add: Word.target_size Word.source_size) lemma of_int_sint_scast: "of_int (sint (x :: 'a :: len word)) = (scast x :: 'b :: len word)" by (metis scast_def word_of_int) lemma scast_of_nat_to_signed [simp]: "scast (of_nat x :: 'a :: len word) = (of_nat x :: 'a signed word)" by (metis cast_simps(23) scast_scast_id(2)) lemma scast_of_nat_signed_to_unsigned_add: "scast (of_nat x + of_nat y :: 'a :: len signed word) = (of_nat x + of_nat y :: 'a :: len word)" by (metis of_nat_add scast_of_nat) lemma scast_of_nat_unsigned_to_signed_add: "(scast (of_nat x + of_nat y :: 'a :: len word)) = (of_nat x + of_nat y :: 'a :: len signed word)" by (metis Abs_fnat_hom_add scast_of_nat_to_signed) lemma and_mask_cases: fixes x :: "'a :: len word" assumes len: "n < LENGTH('a)" shows "x && mask n \ of_nat ` set [0 ..< 2 ^ n]" proof - have "x && mask n \ {0 .. 2 ^ n - 1}" by (simp add: mask_def word_and_le1) also have "... = of_nat ` {0 .. 2 ^ n - 1}" apply (rule set_eqI, rule iffI) apply (clarsimp simp: image_iff) apply (rule_tac x="unat x" in bexI; simp) using len apply (simp add: word_le_nat_alt unat_2tp_if unat_minus_one) using len apply (clarsimp simp: word_le_nat_alt unat_2tp_if unat_minus_one) apply (subst unat_of_nat_eq; simp add: nat_le_Suc_less) apply (erule less_le_trans) apply simp done also have "{0::nat .. 2^n - 1} = set [0 ..< 2^n]" by (auto simp: nat_le_Suc_less) finally show ?thesis . qed lemma sint_of_nat_ge_zero: "x < 2 ^ (LENGTH('a) - 1) \ sint (of_nat x :: 'a :: len word) \ 0" by (simp add: Word_Lemmas.of_nat_power not_msb_from_less sint_eq_uint) lemma sint_eq_uint_2pl: "\ (a :: 'a :: len word) < 2 ^ (LENGTH('a) - 1) \ \ sint a = uint a" by (simp add: not_msb_from_less sint_eq_uint word_2p_lem word_size) lemma sint_of_nat_le: "\ b < 2 ^ (LENGTH('a) - 1); a \ b \ \ sint (of_nat a :: 'a :: len word) \ sint (of_nat b :: 'a :: len word)" by (smt Word_Lemmas.of_nat_power diff_less le_less_trans len_gt_0 len_of_numeral_defs(2) nat_power_minus_less of_nat_le_iff sint_eq_uint_2pl uint_nat unat_of_nat_len) lemma int_eq_sint: "x < 2 ^ (LENGTH('a) - 1) \ sint (of_nat x :: 'a :: len word) = int x" by (smt Word_Lemmas.of_nat_power diff_less le_less_trans len_gt_0 len_of_numeral_defs(2) nat_less_le sint_eq_uint_2pl uint_nat unat_lt2p unat_of_nat_len unat_power_lower) lemma sint_ctz: "LENGTH('a) > 2 \ 0 \ sint (of_nat (word_ctz (x :: 'a :: len word)) :: 'a signed word) \ sint (of_nat (word_ctz x) :: 'a signed word) \ int (LENGTH('a))" apply (subgoal_tac "LENGTH('a) < 2 ^ (LENGTH('a) - 1)") apply (rule conjI) apply (metis len_signed order_le_less_trans sint_of_nat_ge_zero word_ctz_le) apply (metis int_eq_sint len_signed sint_of_nat_le word_ctz_le) by (rule small_powers_of_2, simp) lemma pow_sub_less: "\ a + b \ LENGTH('a); unat (x :: 'a :: len word) = 2 ^ a \ \ unat (x * 2 ^ b - 1) < 2 ^ (a + b)" by (metis (mono_tags) eq_or_less_helperD not_less of_nat_numeral power_add semiring_1_class.of_nat_power unat_pow_le_intro word_unat.Rep_inverse) lemma sle_le_2pl: "\ (b :: 'a :: len word) < 2 ^ (LENGTH('a) - 1); a \ b \ \ a <=s b" by (simp add: not_msb_from_less word_sle_msb_le) lemma sless_less_2pl: "\ (b :: 'a :: len word) < 2 ^ (LENGTH('a) - 1); a < b \ \ a > n = w && mask (size w - n)" by (cases "n \ size w"; clarsimp simp: word_and_le2 and_mask shiftl_zero_size) lemma unat_of_nat_word_log2: "LENGTH('a) < 2 ^ LENGTH('b) \ unat (of_nat (word_log2 (n :: 'a :: len word)) :: 'b :: len word) = word_log2 n" by (metis less_trans unat_of_nat_eq word_log2_max word_size) lemma aligned_sub_aligned_simple: "\ is_aligned a n; is_aligned b n \ \ is_aligned (a - b) n" by (simp add: aligned_sub_aligned) lemma minus_one_shift: "- (1 << n) = (-1 << n :: 'a::len word)" by (simp add: mask_def NOT_eq flip: mul_not_mask_eq_neg_shiftl) lemma ucast_eq_mask: "(UCAST('a::len \ 'b::len) x = UCAST('a \ 'b) y) = (x && mask LENGTH('b) = y && mask LENGTH('b))" by (rule iffI; word_eqI_solve) context fixes w :: "'a::len word" begin private lemma sbintrunc_uint_ucast: assumes "Suc n = LENGTH('b::len)" shows "sbintrunc n (uint (ucast w :: 'b word)) = sbintrunc n (uint w)" by (metis assms sbintrunc_bintrunc ucast_def word_ubin.eq_norm) private lemma test_bit_sbintrunc: assumes "i < LENGTH('a)" shows "(word_of_int (sbintrunc n (uint w)) :: 'a word) !! i = (if n < i then w !! n else w !! i)" using assms by (simp add: nth_sbintr) (simp add: test_bit_bin) private lemma test_bit_sbintrunc_ucast: assumes len_a: "i < LENGTH('a)" shows "(word_of_int (sbintrunc (LENGTH('b) - 1) (uint (ucast w :: 'b word))) :: 'a word) !! i = (if LENGTH('b::len) \ i then w !! (LENGTH('b) - 1) else w !! i)" apply (subst sbintrunc_uint_ucast) apply simp apply (subst test_bit_sbintrunc) apply (rule len_a) apply (rule if_cong[OF _ refl refl]) using leD less_linear by fastforce lemma scast_ucast_high_bits: "scast (ucast w :: 'b::len word) = w \ (\ i \ {LENGTH('b) ..< size w}. w !! i = w !! (LENGTH('b) - 1))" unfolding scast_def sint_uint word_size apply (subst word_eq_iff) apply (rule iffI) apply (rule ballI) apply (drule_tac x=i in spec) apply (subst (asm) test_bit_sbintrunc_ucast; simp) apply (rule allI) apply (case_tac "n < LENGTH('a)") apply (subst test_bit_sbintrunc_ucast) apply simp apply (case_tac "n \ LENGTH('b)") apply (drule_tac x=n in bspec) by auto lemma scast_ucast_mask_compare: "scast (ucast w :: 'b::len word) = w \ (w \ mask (LENGTH('b) - 1) \ ~~(mask (LENGTH('b) - 1)) \ w)" apply (clarsimp simp: le_mask_high_bits neg_mask_le_high_bits scast_ucast_high_bits word_size) apply (rule iffI; clarsimp) apply (rename_tac i j; case_tac "i = LENGTH('b) - 1"; case_tac "j = LENGTH('b) - 1") by auto lemma ucast_less_shiftl_helper': "\ LENGTH('b) + (a::nat) < LENGTH('a); 2 ^ (LENGTH('b) + a) \ n\ \ (ucast (x :: 'b::len word) << a) < (n :: 'a::len word)" apply (erule order_less_le_trans[rotated]) using ucast_less[where x=x and 'a='a] apply (simp only: shiftl_t2n field_simps) apply (rule word_less_power_trans2; simp) done end lemma ucast_ucast_mask2: "is_down (UCAST ('a \ 'b)) \ UCAST ('b::len \ 'c::len) (UCAST ('a::len \ 'b::len) x) = UCAST ('a \ 'c) (x && mask LENGTH('b))" by word_eqI_solve lemma ucast_NOT: "ucast (~~x) = ~~(ucast x) && mask (LENGTH('a))" for x::"'a::len word" by word_eqI lemma ucast_NOT_down: "is_down UCAST('a::len \ 'b::len) \ UCAST('a \ 'b) (~~x) = ~~(UCAST('a \ 'b) x)" by word_eqI lemma of_bl_mult_and_not_mask_eq: "\is_aligned (a :: 'a::len word) n; length b + m \ n\ \ a + of_bl b * (2^m) && ~~(mask n) = a" by (smt add.left_neutral add_diff_cancel_right' add_mask_lower_bits and_mask_plus is_aligned_mask is_aligned_weaken le_less_trans of_bl_length2 subtract_mask(1)) lemma bin_to_bl_of_bl_eq: "\is_aligned (a::'a::len word) n; length b + c \ n; length b + c < LENGTH('a)\ \ bin_to_bl (length b) (uint ((a + of_bl b * 2^c) >> c)) = b" apply (subst word_plus_and_or_coroll) apply (erule is_aligned_get_word_bits) apply (rule is_aligned_AND_less_0) apply (simp add: is_aligned_mask) apply (rule order_less_le_trans) apply (rule of_bl_length2) apply simp apply (simp add: two_power_increasing) apply simp apply (rule nth_equalityI) apply (simp only: len_bin_to_bl) apply (clarsimp simp only: len_bin_to_bl nth_bin_to_bl word_test_bit_def[symmetric]) apply (simp add: nth_shiftr nth_shiftl shiftl_t2n[where n=c, simplified mult.commute, simplified, symmetric]) apply (simp add: is_aligned_nth[THEN iffD1, rule_format] test_bit_of_bl nth_rev) apply arith done end diff --git a/thys/Word_Lib/Word_Lib.thy b/thys/Word_Lib/Word_Lib.thy --- a/thys/Word_Lib/Word_Lib.thy +++ b/thys/Word_Lib/Word_Lib.thy @@ -1,694 +1,693 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Additional Word Operations" theory Word_Lib imports Word_Syntax begin definition ptr_add :: "'a :: len word \ nat \ 'a word" where "ptr_add ptr n \ ptr + of_nat n" definition complement :: "'a :: len word \ 'a word" where "complement x \ ~~ x" definition alignUp :: "'a::len word \ nat \ 'a word" where "alignUp x n \ x + 2 ^ n - 1 && complement (2 ^ n - 1)" (* standard notation for blocks of 2^n-1 words, usually aligned; abbreviation so it simplifies directly *) abbreviation mask_range :: "'a::len word \ nat \ 'a word set" where "mask_range p n \ {p .. p + mask n}" (* Haskellish names/syntax *) notation (input) test_bit ("testBit") definition w2byte :: "'a :: len word \ 8 word" where "w2byte \ ucast" (* * Signed division: when the result of a division is negative, * we will round towards zero instead of towards minus infinity. *) class signed_div = fixes sdiv :: "'a \ 'a \ 'a" (infixl "sdiv" 70) fixes smod :: "'a \ 'a \ 'a" (infixl "smod" 70) instantiation int :: signed_div begin definition "(a :: int) sdiv b \ sgn (a * b) * (abs a div abs b)" definition "(a :: int) smod b \ a - (a sdiv b) * b" instance .. end instantiation word :: (len) signed_div begin definition "(a :: ('a::len) word) sdiv b = word_of_int (sint a sdiv sint b)" definition "(a :: ('a::len) word) smod b = word_of_int (sint a smod sint b)" instance .. end (* Tests *) lemma "( 4 :: word32) sdiv 4 = 1" "(-4 :: word32) sdiv 4 = -1" "(-3 :: word32) sdiv 4 = 0" "( 3 :: word32) sdiv -4 = 0" "(-3 :: word32) sdiv -4 = 0" "(-5 :: word32) sdiv -4 = 1" "( 5 :: word32) sdiv -4 = -1" by (simp_all add: sdiv_word_def sdiv_int_def) lemma "( 4 :: word32) smod 4 = 0" "( 3 :: word32) smod 4 = 3" "(-3 :: word32) smod 4 = -3" "( 3 :: word32) smod -4 = 3" "(-3 :: word32) smod -4 = -3" "(-5 :: word32) smod -4 = -1" "( 5 :: word32) smod -4 = 1" by (simp_all add: smod_word_def smod_int_def sdiv_int_def) (* Count leading zeros *) definition word_clz :: "'a::len word \ nat" where "word_clz w \ length (takeWhile Not (to_bl w))" (* Count trailing zeros *) definition word_ctz :: "'a::len word \ nat" where "word_ctz w \ length (takeWhile Not (rev (to_bl w)))" definition word_log2 :: "'a::len word \ nat" where "word_log2 (w::'a::len word) \ size w - 1 - word_clz w" (* Bit population count. Equivalent of __builtin_popcount. *) definition pop_count :: "('a::len) word \ nat" where "pop_count w \ length (filter id (to_bl w))" (* Sign extension from bit n *) definition sign_extend :: "nat \ 'a::len word \ 'a word" where "sign_extend n w \ if w !! n then w || ~~ (mask n) else w && mask n" definition sign_extended :: "nat \ 'a::len word \ bool" where "sign_extended n w \ \i. n < i \ i < size w \ w !! i = w !! n" lemma ptr_add_0 [simp]: "ptr_add ref 0 = ref " unfolding ptr_add_def by simp lemma shiftl_power: "(shiftl1 ^^ x) (y::'a::len word) = 2 ^ x * y" apply (induct x) apply simp apply (simp add: shiftl1_2t) done lemmas of_bl_reasoning = to_bl_use_of_bl of_bl_append lemma uint_of_bl_is_bl_to_bin_drop: "length (dropWhile Not l) \ LENGTH('a) \ uint (of_bl l :: 'a::len word) = bl_to_bin l" apply (simp add: of_bl_def) apply (rule word_uint.Abs_inverse) apply (simp add: uints_num bl_to_bin_ge0) apply (rule order_less_le_trans) apply (rule bl_to_bin_lt2p_drop) apply (simp) done corollary uint_of_bl_is_bl_to_bin: "length l\LENGTH('a) \ uint ((of_bl::bool list\ ('a :: len) word) l) = bl_to_bin l" apply(rule uint_of_bl_is_bl_to_bin_drop) using le_trans length_dropWhile_le by blast lemma bin_to_bl_or: "bin_to_bl n (a OR b) = map2 (\) (bin_to_bl n a) (bin_to_bl n b)" using bl_or_aux_bin[where n=n and v=a and w=b and bs="[]" and cs="[]"] by simp lemma word_ops_nth [simp]: shows word_or_nth: "(x || y) !! n = (x !! n \ y !! n)" and word_and_nth: "(x && y) !! n = (x !! n \ y !! n)" and word_xor_nth: "(x xor y) !! n = (x !! n \ y !! n)" by ((cases "n < size x", auto dest: test_bit_size simp: word_ops_nth_size word_size)[1])+ (* simp del to avoid warning on the simp add in iff *) declare test_bit_1 [simp del, iff] (* test: *) lemma "1 < (1024::32 word) \ 1 \ (1024::32 word)" by simp lemma and_not_mask: "w AND NOT (mask n) = (w >> n) << n" apply (rule word_eqI) apply (simp add : word_ops_nth_size word_size) apply (simp add : nth_shiftr nth_shiftl) by auto lemma and_mask: "w AND mask n = (w << (size w - n)) >> (size w - n)" apply (rule word_eqI) apply (simp add : word_ops_nth_size word_size) apply (simp add : nth_shiftr nth_shiftl) by auto lemma AND_twice [simp]: "(w && m) && m = w && m" by (simp add: word_eqI) lemma word_combine_masks: "w && m = z \ w && m' = z' \ w && (m || m') = (z || z')" by (auto simp: word_eq_iff) lemma nth_w2p_same: "(2^n :: 'a :: len word) !! n = (n < LENGTH('a))" by (simp add : nth_w2p) lemma p2_gt_0: "(0 < (2 ^ n :: 'a :: len word)) = (n < LENGTH('a))" apply (simp add : word_gt_0) apply safe apply (erule swap) apply (rule word_eqI) apply (simp add : nth_w2p) apply (drule word_eqD) apply simp apply (erule notE) apply (erule nth_w2p_same [THEN iffD2]) done lemmas uint_2p_alt = uint_2p [unfolded p2_gt_0] lemma shiftr_div_2n_w: "n < size w \ w >> n = w div (2^n :: 'a :: len word)" apply (unfold word_div_def) apply (simp add : uint_2p_alt word_size) apply (rule word_uint.Rep_inverse' [THEN sym]) apply (rule shiftr_div_2n) done lemmas less_def = less_eq [symmetric] lemmas le_def = not_less [symmetric, where ?'a = nat] lemmas p2_eq_0 [simp] = trans [OF eq_commute iffD2 [OF Not_eq_iff p2_gt_0, folded le_def, unfolded word_gt_0 not_not]] lemma neg_mask_is_div': "n < size w \ w AND NOT (mask n) = ((w div (2 ^ n)) * (2 ^ n))" by (simp add : and_not_mask shiftr_div_2n_w shiftl_t2n word_size) lemma neg_mask_is_div: "w AND NOT (mask n) = (w div 2^n) * 2^n" apply (cases "n < size w") apply (erule neg_mask_is_div') apply (simp add: word_size) apply (frule p2_gt_0 [THEN Not_eq_iff [THEN iffD2], THEN iffD2]) apply (simp add: word_gt_0 del: p2_eq_0) apply (rule word_eqI) apply (simp add: word_ops_nth_size word_size) done lemma and_mask_arith': "0 < n \ w AND mask n = (w * 2^(size w - n)) div 2^(size w - n)" by (simp add: and_mask shiftr_div_2n_w shiftl_t2n word_size mult.commute) lemmas p2len = iffD2 [OF p2_eq_0 order_refl] lemma and_mask_arith: "w AND mask n = (w * 2^(size w - n)) div 2^(size w - n)" apply (cases "0 < n") - apply (erule and_mask_arith') + apply (auto elim!: and_mask_arith') apply (simp add: word_size) - apply (simp add: word_of_int_hom_syms word_div_def) done lemma mask_2pm1: "mask n = 2 ^ n - 1" by (simp add : mask_def) lemma add_mask_fold: "x + 2 ^ n - 1 = x + mask n" by (simp add: mask_def) lemma word_and_mask_le_2pm1: "w && mask n \ 2 ^ n - 1" by (simp add: mask_2pm1[symmetric] word_and_le1) lemma is_aligned_AND_less_0: "u && mask n = 0 \ v < 2^n \ u && v = 0" apply (drule less_mask_eq) apply (simp add: mask_2pm1) apply (rule word_eqI) apply (clarsimp simp add: word_size) apply (drule_tac x=na in word_eqD) apply (drule_tac x=na in word_eqD) apply simp done lemma le_shiftr1: "u <= v \ shiftr1 u <= shiftr1 v" apply (unfold word_le_def shiftr1_def word_ubin.eq_norm) apply (unfold bin_rest_trunc_i trans [OF bintrunc_bintrunc_l word_ubin.norm_Rep, unfolded word_ubin.norm_Rep, OF order_refl [THEN le_SucI]]) apply (case_tac "uint u" rule: bin_exhaust) apply (rename_tac bs bit) apply (case_tac "uint v" rule: bin_exhaust) apply (rename_tac bs' bit') apply (case_tac "bit") apply (case_tac "bit'", auto simp: less_eq_int_code le_Bits)[1] apply (case_tac bit') apply (simp add: le_Bits less_eq_int_code) apply (auto simp: le_Bits less_eq_int_code) done lemma le_shiftr: - "u \ v \ u >> (n :: nat) \ (v :: 'a :: len0 word) >> n" + "u \ v \ u >> (n :: nat) \ (v :: 'a :: len word) >> n" apply (unfold shiftr_def) apply (induct_tac "n") apply auto apply (erule le_shiftr1) done lemma shiftr_mask_le: "n <= m \ mask n >> m = 0" apply (rule word_eqI) apply (simp add: word_size nth_shiftr) done lemmas shiftr_mask = order_refl [THEN shiftr_mask_le, simp] lemma word_leI: - "(\n. \n < size (u::'a::len0 word); u !! n \ \ (v::'a::len0 word) !! n) \ u <= v" + "(\n. \n < size (u::'a::len word); u !! n \ \ (v::'a::len word) !! n) \ u <= v" apply (rule xtr4) apply (rule word_and_le2) apply (rule word_eqI) apply (simp add: word_ao_nth) apply safe apply assumption apply (erule_tac [2] asm_rl) apply (unfold word_size) by auto lemma le_mask_iff: "(w \ mask n) = (w >> n = 0)" apply safe apply (rule word_le_0_iff [THEN iffD1]) apply (rule xtr3) apply (erule_tac [2] le_shiftr) apply simp apply (rule word_leI) apply (rename_tac n') apply (drule_tac x = "n' - n" in word_eqD) apply (simp add : nth_shiftr word_size) apply (case_tac "n <= n'") by auto lemma and_mask_eq_iff_shiftr_0: "(w AND mask n = w) = (w >> n = 0)" apply (unfold test_bit_eq_iff [THEN sym]) apply (rule iffI) apply (rule ext) apply (rule_tac [2] ext) apply (auto simp add : word_ao_nth nth_shiftr) apply (drule arg_cong) apply (drule iffD2) apply assumption apply (simp add : word_ao_nth) prefer 2 apply (simp add : word_size test_bit_bin) apply (drule_tac f = "%u. u !! (x - n)" in arg_cong) apply (simp add : nth_shiftr) apply (case_tac "n <= x") apply auto done lemmas and_mask_eq_iff_le_mask = trans [OF and_mask_eq_iff_shiftr_0 le_mask_iff [THEN sym]] lemma mask_shiftl_decompose: "mask m << n = mask (m + n) && ~~ (mask n)" by (auto intro!: word_eqI simp: and_not_mask nth_shiftl nth_shiftr word_size) lemma one_bit_shiftl: "set_bit 0 n True = (1 :: 'a :: len word) << n" apply (rule word_eqI) apply (auto simp add: test_bit_set_gen nth_shiftl word_size simp del: word_set_bit_0 shiftl_1) done lemmas one_bit_pow = trans [OF one_bit_shiftl shiftl_1] lemmas bin_sc_minus_simps = bin_sc_simps (2,3,4) [THEN [2] trans, OF bin_sc_minus [THEN sym]] lemma NOT_eq: "NOT (x :: 'a :: len word) = - x - 1" apply (cut_tac x = "x" in word_add_not) apply (drule add.commute [THEN trans]) apply (drule eq_diff_eq [THEN iffD2]) by simp lemma NOT_mask: "NOT (mask n) = -(2 ^ n)" by (simp add : NOT_eq mask_2pm1) lemma le_m1_iff_lt: "(x > (0 :: 'a :: len word)) = ((y \ x - 1) = (y < x))" by uint_arith lemmas gt0_iff_gem1 = iffD1 [OF iffD1 [OF iff_left_commute le_m1_iff_lt] order_refl] lemmas power_2_ge_iff = trans [OF gt0_iff_gem1 [THEN sym] p2_gt_0] lemma le_mask_iff_lt_2n: "n < len_of TYPE ('a) = (((w :: 'a :: len word) \ mask n) = (w < 2 ^ n))" unfolding mask_2pm1 by (rule trans [OF p2_gt_0 [THEN sym] le_m1_iff_lt]) lemmas mask_lt_2pn = le_mask_iff_lt_2n [THEN iffD1, THEN iffD1, OF _ order_refl] lemma bang_eq: - fixes x :: "'a::len0 word" + fixes x :: "'a::len word" shows "(x = y) = (\n. x !! n = y !! n)" by (subst test_bit_eq_iff[symmetric]) fastforce lemma word_unat_power: "(2 :: 'a :: len word) ^ n = of_nat (2 ^ n)" by simp lemma of_nat_mono_maybe: assumes xlt: "x < 2 ^ len_of TYPE ('a)" shows "y < x \ of_nat y < (of_nat x :: 'a :: len word)" apply (subst word_less_nat_alt) apply (subst unat_of_nat)+ apply (subst mod_less) apply (erule order_less_trans [OF _ xlt]) apply (subst mod_less [OF xlt]) apply assumption done lemma shiftl_over_and_dist: fixes a::"'a::len word" shows "(a AND b) << c = (a << c) AND (b << c)" apply(rule word_eqI) apply(simp add: word_ao_nth nth_shiftl, safe) done lemma shiftr_over_and_dist: fixes a::"'a::len word" shows "a AND b >> c = (a >> c) AND (b >> c)" apply(rule word_eqI) apply(simp add:nth_shiftr word_ao_nth) done lemma sshiftr_over_and_dist: fixes a::"'a::len word" shows "a AND b >>> c = (a >>> c) AND (b >>> c)" apply(rule word_eqI) apply(simp add:nth_sshiftr word_ao_nth word_size) done lemma shiftl_over_or_dist: fixes a::"'a::len word" shows "a OR b << c = (a << c) OR (b << c)" apply(rule word_eqI) apply(simp add:nth_shiftl word_ao_nth, safe) done lemma shiftr_over_or_dist: fixes a::"'a::len word" shows "a OR b >> c = (a >> c) OR (b >> c)" apply(rule word_eqI) apply(simp add:nth_shiftr word_ao_nth) done lemma sshiftr_over_or_dist: fixes a::"'a::len word" shows "a OR b >>> c = (a >>> c) OR (b >>> c)" apply(rule word_eqI) apply(simp add:nth_sshiftr word_ao_nth word_size) done lemmas shift_over_ao_dists = shiftl_over_or_dist shiftr_over_or_dist sshiftr_over_or_dist shiftl_over_and_dist shiftr_over_and_dist sshiftr_over_and_dist lemma shiftl_shiftl: fixes a::"'a::len word" shows "a << b << c = a << (b + c)" apply(rule word_eqI) apply(auto simp:word_size nth_shiftl add.commute add.left_commute) done lemma shiftr_shiftr: fixes a::"'a::len word" shows "a >> b >> c = a >> (b + c)" apply(rule word_eqI) apply(simp add:word_size nth_shiftr add.left_commute add.commute) done lemma shiftl_shiftr1: fixes a::"'a::len word" shows "c \ b \ a << b >> c = a AND (mask (size a - b)) << (b - c)" apply(rule word_eqI) apply(auto simp:nth_shiftr nth_shiftl word_size word_ao_nth) done lemma shiftl_shiftr2: fixes a::"'a::len word" shows "b < c \ a << b >> c = (a >> (c - b)) AND (mask (size a - c))" apply(rule word_eqI) apply(auto simp:nth_shiftr nth_shiftl word_size word_ao_nth) done lemma shiftr_shiftl1: fixes a::"'a::len word" shows "c \ b \ a >> b << c = (a >> (b - c)) AND (NOT (mask c))" apply(rule word_eqI) apply(auto simp:nth_shiftr nth_shiftl word_size word_ops_nth_size) done lemma shiftr_shiftl2: fixes a::"'a::len word" shows "b < c \ a >> b << c = (a << (c - b)) AND (NOT (mask c))" apply(rule word_eqI) apply(auto simp:nth_shiftr nth_shiftl word_size word_ops_nth_size) done lemmas multi_shift_simps = shiftl_shiftl shiftr_shiftr shiftl_shiftr1 shiftl_shiftr2 shiftr_shiftl1 shiftr_shiftl2 lemma word_and_max_word: fixes a::"'a::len word" shows "x = max_word \ a AND x = a" by simp lemma word_and_1_bl: fixes x::"'a::len word" shows "(x AND 1) = of_bl [x !! 0]" by (simp add: word_and_1) lemma word_1_and_bl: fixes x::"'a::len word" shows "(1 AND x) = of_bl [x !! 0]" by (subst word_bw_comms) (simp add: word_and_1) lemma scast_scast_id [simp]: "scast (scast x :: ('a::len) signed word) = (x :: 'a word)" "scast (scast y :: ('a::len) word) = (y :: 'a signed word)" by (auto simp: is_up scast_up_scast_id) lemma scast_ucast_id [simp]: "scast (ucast (x :: 'a::len word) :: 'a signed word) = x" by (metis down_cast_same is_down len_signed order_refl scast_scast_id(1)) lemma ucast_scast_id [simp]: "ucast (scast (x :: 'a::len signed word) :: 'a word) = x" by (metis scast_scast_id(2) scast_ucast_id) lemma scast_of_nat [simp]: "scast (of_nat x :: 'a::len signed word) = (of_nat x :: 'a word)" by (metis (hide_lams, no_types) len_signed scast_def uint_sint word_of_nat word_ubin.Abs_norm word_ubin.eq_norm) lemma ucast_of_nat: "is_down (ucast :: 'a :: len word \ 'b :: len word) \ ucast (of_nat n :: 'a word) = (of_nat n :: 'b word)" apply (rule sym) apply (subst word_unat.inverse_norm) apply (simp add: ucast_def word_of_int[symmetric] of_nat_nat[symmetric] unat_def[symmetric]) apply (simp add: unat_of_nat) apply (rule nat_int.Rep_eqD) apply (simp only: zmod_int) apply (rule mod_mod_cancel) by (simp add: is_down le_imp_power_dvd) (* shortcut for some specific lengths *) lemma word_fixed_sint_1[simp]: "sint (1::8 word) = 1" "sint (1::16 word) = 1" "sint (1::32 word) = 1" "sint (1::64 word) = 1" by (auto simp: sint_word_ariths) lemma word_sint_1 [simp]: "sint (1::'a::len word) = (if LENGTH('a) = 1 then -1 else 1)" apply (case_tac "LENGTH('a) \ 1") apply (metis diff_is_0_eq sbintrunc_0_numeral(1) sint_n1 word_1_wi word_m1_wi word_msb_1 word_msb_n1 word_sbin.Abs_norm) apply (metis bin_nth_1 diff_is_0_eq neq0_conv sbintrunc_minus_simps(4) sint_word_ariths(8) uint_1 word_msb_1 word_msb_nth) done lemma scast_1': "(scast (1::'a::len word) :: 'b::len word) = (word_of_int (sbintrunc (LENGTH('a::len) - Suc 0) (1::int)))" by (metis One_nat_def scast_def sint_word_ariths(8)) lemma scast_1 [simp]: "(scast (1::'a::len word) :: 'b::len word) = (if LENGTH('a) = 1 then -1 else 1)" by (clarsimp simp: scast_1') (metis Suc_pred len_gt_0 nat.exhaust sbintrunc_Suc_numeral(1) uint_1 word_uint.Rep_inverse') lemma scast_eq_scast_id [simp]: "((scast (a :: 'a::len signed word) :: 'a word) = scast b) = (a = b)" by (metis ucast_scast_id) lemma ucast_eq_ucast_id [simp]: "((ucast (a :: 'a::len word) :: 'a signed word) = ucast b) = (a = b)" by (metis scast_ucast_id) lemma scast_ucast_norm [simp]: "(ucast (a :: 'a::len word) = (b :: 'a signed word)) = (a = scast b)" "((b :: 'a signed word) = ucast (a :: 'a::len word)) = (a = scast b)" by (metis scast_ucast_id ucast_scast_id)+ lemma of_bl_drop: "of_bl (drop n xs) = (of_bl xs && mask (length xs - n))" apply (clarsimp simp: bang_eq test_bit_of_bl rev_nth cong: rev_conj_cong) apply (safe; simp add: word_size to_bl_nth) done lemma of_int_uint [simp]: "of_int (uint x) = x" by (metis word_of_int word_uint.Rep_inverse') lemma shiftr_mask2: "n \ LENGTH('a) \ (mask n >> m :: ('a :: len) word) = mask (n - m)" apply (rule word_eqI) apply (simp add: nth_shiftr word_size) apply arith done corollary word_plus_and_or_coroll: "x && y = 0 \ x + y = x || y" using word_plus_and_or[where x=x and y=y] by simp corollary word_plus_and_or_coroll2: "(x && w) + (x && ~~ w) = x" apply (subst word_plus_and_or_coroll) apply (rule word_eqI, simp add: word_size word_ops_nth_size) apply (rule word_eqI, simp add: word_size word_ops_nth_size) apply blast done lemma less_le_mult_nat': "w * c < b * c ==> 0 \ c ==> Suc w * c \ b * (c::nat)" apply (rule mult_right_mono) apply (rule Suc_leI) apply (erule (1) mult_right_less_imp_less) apply assumption done lemmas less_le_mult_nat = less_le_mult_nat'[simplified distrib_right, simplified] (* FIXME: these should eventually be moved to HOL/Word. *) lemmas extra_sle_sless_unfolds [simp] = word_sle_def[where a=0 and b=1] word_sle_def[where a=0 and b="numeral n"] word_sle_def[where a=1 and b=0] word_sle_def[where a=1 and b="numeral n"] word_sle_def[where a="numeral n" and b=0] word_sle_def[where a="numeral n" and b=1] word_sless_alt[where a=0 and b=1] word_sless_alt[where a=0 and b="numeral n"] word_sless_alt[where a=1 and b=0] word_sless_alt[where a=1 and b="numeral n"] word_sless_alt[where a="numeral n" and b=0] word_sless_alt[where a="numeral n" and b=1] for n lemma to_bl_1: "to_bl (1::'a::len word) = replicate (LENGTH('a) - 1) False @ [True]" proof - have "to_bl (1 :: 'a::len word) = to_bl (mask 1 :: 'a::len word)" by (simp add: mask_def) also have "\ = replicate (LENGTH('a) - 1) False @ [True]" by (cases "LENGTH('a)"; clarsimp simp: to_bl_mask) finally show ?thesis . qed lemma list_of_false: "True \ set xs \ xs = replicate (length xs) False" by (induct xs, simp_all) lemma eq_zero_set_bl: "(w = 0) = (True \ set (to_bl w))" using list_of_false word_bl.Rep_inject by fastforce lemma diff_diff_less: "(i < m - (m - (n :: nat))) = (i < m \ i < n)" by auto lemma pop_count_0[simp]: "pop_count 0 = 0" by (clarsimp simp:pop_count_def) lemma pop_count_1[simp]: "pop_count 1 = 1" by (clarsimp simp:pop_count_def to_bl_1) lemma pop_count_0_imp_0: "(pop_count w = 0) = (w = 0)" apply (rule iffI) apply (clarsimp simp:pop_count_def) apply (subst (asm) filter_empty_conv) apply (clarsimp simp:eq_zero_set_bl) apply fast apply simp done end diff --git a/thys/Word_Lib/Word_Next.thy b/thys/Word_Lib/Word_Next.thy --- a/thys/Word_Lib/Word_Next.thy +++ b/thys/Word_Lib/Word_Next.thy @@ -1,95 +1,95 @@ (* SPDX-License-Identifier: BSD-3-Clause *) section\Increment and Decrement Machine Words Without Wrap-Around\ theory Word_Next imports Aligned begin text\Previous and next words addresses, without wrap around.\ definition word_next :: "'a::len word \ 'a::len word" where "word_next a \ if a = max_word then max_word else a + 1" definition word_prev :: "'a::len word \ 'a::len word" where "word_prev a \ if a = 0 then 0 else a - 1" text\Examples:\ lemma "word_next (2:: 8 word) = 3" by eval lemma "word_next (255:: 8 word) = 255" by eval lemma "word_prev (2:: 8 word) = 1" by eval lemma "word_prev (0:: 8 word) = 0" by eval lemma plus_one_helper[elim!]: "x < n + (1 :: 'a :: len word) \ x \ n" apply (simp add: word_less_nat_alt word_le_nat_alt field_simps) apply (case_tac "1 + n = 0") apply simp apply (subst(asm) unatSuc, assumption) apply arith done lemma plus_one_helper2: "\ x \ n; n + 1 \ 0 \ \ x < n + (1 :: 'a :: len word)" by (simp add: word_less_nat_alt word_le_nat_alt field_simps unatSuc) lemma less_x_plus_1: fixes x :: "'a :: len word" shows "x \ max_word \ (y < (x + 1)) = (y < x \ y = x)" apply (rule iffI) apply (rule disjCI) apply (drule plus_one_helper) apply simp apply (subgoal_tac "x < x + 1") apply (erule disjE, simp_all) apply (rule plus_one_helper2 [OF order_refl]) apply (rule notI, drule max_word_wrap) apply simp done lemma word_Suc_leq: fixes k::"'a::len word" shows "k \ max_word \ x < k + 1 \ x \ k" using less_x_plus_1 word_le_less_eq by auto lemma word_Suc_le: fixes k::"'a::len word" shows "x \ max_word \ x + 1 \ k \ x < k" by (meson not_less word_Suc_leq) lemma word_lessThan_Suc_atMost: \{..< k + 1} = {..k}\ if \k \ - 1\ for k :: \'a::len word\ using that by (simp add: lessThan_def atMost_def word_Suc_leq) lemma word_atLeastLessThan_Suc_atLeastAtMost: \{l ..< u + 1} = {l..u}\ if \u \ - 1\ for l :: \'a::len word\ using that by (simp add: atLeastAtMost_def atLeastLessThan_def word_lessThan_Suc_atMost) lemma word_atLeastAtMost_Suc_greaterThanAtMost: \{m<..u} = {m + 1..u}\ if \m \ - 1\ for m :: \'a::len word\ using that by (simp add: greaterThanAtMost_def greaterThan_def atLeastAtMost_def atLeast_def word_Suc_le) lemma word_atLeastLessThan_Suc_atLeastAtMost_union: fixes l::"'a::len word" assumes "m \ max_word" and "l \ m" and "m \ u" shows "{l..m} \ {m+1..u} = {l..u}" proof - from ivl_disj_un_two(8)[OF assms(2) assms(3)] have "{l..u} = {l..m} \ {m<..u}" by blast with assms show ?thesis by(simp add: word_atLeastAtMost_Suc_greaterThanAtMost) qed lemma max_word_less_eq_iff [simp]: - \- 1 \ w \ w = - 1\ for w :: \'a::len0 word\ + \- 1 \ w \ w = - 1\ for w :: \'a::len word\ by (auto simp add: le_less) lemma word_adjacent_union: "word_next e = s' \ s \ e \ s' \ e' \ {s..e} \ {s'..e'} = {s .. e'}" apply (simp add: word_next_def ivl_disj_un_two_touch split: if_splits) apply (drule sym) apply simp apply (subst word_atLeastLessThan_Suc_atLeastAtMost_union) apply (simp_all add: word_Suc_le) done end diff --git a/thys/Word_Lib/Word_Syntax.thy b/thys/Word_Lib/Word_Syntax.thy --- a/thys/Word_Lib/Word_Syntax.thy +++ b/thys/Word_Lib/Word_Syntax.thy @@ -1,64 +1,64 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Additional Syntax for Word Bit Operations" theory Word_Syntax imports "HOL-Word.More_Word" WordBitwise_Signed Hex_Words Norm_Words Word_Type_Syntax begin text \Additional bit and type syntax that forces word types.\ type_synonym word8 = "8 word" type_synonym word16 = "16 word" type_synonym word32 = "32 word" type_synonym word64 = "64 word" lemma len8: "len_of (x :: 8 itself) = 8" by simp lemma len16: "len_of (x :: 16 itself) = 16" by simp lemma len32: "len_of (x :: 32 itself) = 32" by simp lemma len64: "len_of (x :: 64 itself) = 64" by simp abbreviation - wordNOT :: "'a::len0 word \ 'a word" ("~~ _" [70] 71) + wordNOT :: "'a::len word \ 'a word" ("~~ _" [70] 71) where "~~ x == NOT x" abbreviation - wordAND :: "'a::len0 word \ 'a word \ 'a word" (infixr "&&" 64) + wordAND :: "'a::len word \ 'a word \ 'a word" (infixr "&&" 64) where "a && b == a AND b" abbreviation - wordOR :: "'a::len0 word \ 'a word \ 'a word" (infixr "||" 59) + wordOR :: "'a::len word \ 'a word \ 'a word" (infixr "||" 59) where "a || b == a OR b" abbreviation - wordXOR :: "'a::len0 word \ 'a word \ 'a word" (infixr "xor" 59) + wordXOR :: "'a::len word \ 'a word \ 'a word" (infixr "xor" 59) where "a xor b == a XOR b" (* testing for presence of word_bitwise *) lemma "((x :: word32) >> 3) AND 7 = (x AND 56) >> 3" by word_bitwise (* FIXME: move to Word distribution *) lemma bin_nth_minus_Bit0[simp]: "0 < n \ bin_nth (numeral (num.Bit0 w)) n = bin_nth (numeral w) (n - 1)" by (cases n; simp) lemma bin_nth_minus_Bit1[simp]: "0 < n \ bin_nth (numeral (num.Bit1 w)) n = bin_nth (numeral w) (n - 1)" by (cases n; simp) end