diff --git a/thys/IMP2_Binary_Heap/IMP2_Binary_Heap.thy b/thys/IMP2_Binary_Heap/IMP2_Binary_Heap.thy --- a/thys/IMP2_Binary_Heap/IMP2_Binary_Heap.thy +++ b/thys/IMP2_Binary_Heap/IMP2_Binary_Heap.thy @@ -1,903 +1,903 @@ theory IMP2_Binary_Heap imports IMP2.IMP2 IMP2.IMP2_Aux_Lemmas begin section \Introduction\ text \In this submission imperative versions of the following array-based binary minimum heap functions are implemented and verified: insert, get-min, delete-min, make-heap. The latter three are then used to prove the correctness of an in-place heapsort, which sorts an array in descending order. To do that in Isabelle/HOL, the proof framework IMP2 \cite{IMP2-AFP} is used. Here arrays are modeled by \int \ int\ functions. The imperative implementations are iterative versions of the partly recursive algorithms described in \cite{MS} and \cite{CLRS}. This submission starts with the basic definitions and lemmas, which are needed for array-based binary heaps. These definitions and lemmas are parameterised with an arbitrary (transitive) comparison function (where such a function is needed), so they are not only applicable to minimum heaps. After some more general, useful lemmas on arrays, the imperative minimum heap functions and the heapsort are implemented and verified.\ section \Heap Related Definitions and Theorems\ subsection \Array Bounds\ text \A small helper function is used to define valid array indices. Note that the lower index bound \l\ is arbitrary and not fixed to 0 or 1. The upper index bound \r\ is not a valid index itself, so that the empty array can be denoted by having \l = r\.\ abbreviation bounded :: "int \ int \ int \ bool" where "bounded l r x \ l \ x \ x < r" subsection \Parent and Children\ subsubsection \Definitions\ text \For the notion of an array-based binary heap, the parent and child relations on the array indices need to be defined.\ definition parent :: "int \ int \ int" where "parent l c = l + (c - l - 1) div 2" definition l_child :: "int \ int \ int" where "l_child l p = 2 * p - l + 1" definition r_child :: "int \ int \ int" where "r_child l p = 2 * p - l + 2" subsubsection \Lemmas\ lemma parent_upper_bound: "parent l c < c \ l \ c" unfolding parent_def by auto lemma parent_upper_bound_alt: "l \ parent l c \ parent l c < c" unfolding parent_def by simp lemma parent_lower_bound: "l \ parent l c \ l < c" unfolding parent_def by linarith lemma grand_parent_upper_bound: "parent l (parent l c) < c \ l \ c" unfolding parent_def by linarith corollary parent_bounds: "l < x \ x < r \ bounded l r (parent l x)" using parent_lower_bound parent_upper_bound_alt by fastforce lemma l_child_lower_bound: " p < l_child l p \ l \ p" unfolding l_child_def by simp corollary l_child_lower_bound_alt: "l \ x \ x \ p \ x < l_child l p" using l_child_lower_bound[of p l] by linarith lemma parent_l_child[simp]: "parent l (l_child l n) = n" unfolding parent_def l_child_def by simp lemma r_child_lower_bound: "l \ p \ p < r_child l p" unfolding r_child_def by simp corollary r_child_lower_bound_alt: "l \ x \ x \ p \ x < r_child l p" using r_child_lower_bound[of l p] by linarith lemma parent_r_child[simp]: "parent l (r_child l n) = n" unfolding parent_def r_child_def by simp lemma smaller_l_child: "l_child l x < r_child l x" unfolding l_child_def r_child_def by simp lemma parent_two_children: "(c = l_child l p \ c = r_child l p) \ parent l c = p" unfolding parent_def l_child_def r_child_def by linarith subsection \Heap Invariants\ subsubsection \Definitions\ text \The following heap invariants and the following lemmas are parameterised with an arbitrary (transitive) comparison function. For the concrete function implementations at the end of this submission \\\ on ints is used.\ text \For the \make_heap\ function, which transforms an unordered array into a valid heap, the notion of a partial heap is needed. Here the heap invariant only holds for array indices between a certain valid array index \m\ and \r\. The standard heap invariant is then simply the special case where \m = l\.\ definition is_partial_heap :: "('a::order \ 'a::order \ bool) \ (int \ 'a::order) \ int \ int \ int \ bool" where "is_partial_heap cmp heap l m r = (\ x. bounded m r x \ bounded m r (parent l x) \ cmp (heap (parent l x)) (heap x))" abbreviation is_heap :: "('a::order \ 'a::order \ bool) \ (int \ 'a::order) \ int \ int \ bool" where "is_heap cmp heap l r \ is_partial_heap cmp heap l l r" text \During all of the modifying heap functions the heap invariant is temporarily violated at a single index \i\ and it is then gradually restored by either \sift_down\ or \sift_up\. The following definitions formalize these weakened invariants. The second part of the conjunction in the following definitions states, that the comparison between the parent of \i\ and each of the children of \i\ evaluates to \True\ without explicitly using the child relations.\ definition is_partial_heap_except_down :: "('a::order \ 'a::order \ bool) \ (int \ 'a::order) \ int \ int \ int \ int \ bool" where "is_partial_heap_except_down cmp heap l m r i = (\ x. bounded m r x \ ((parent l x \ i \ bounded m r (parent l x) \ cmp (heap (parent l x)) (heap x)) \ (parent l x = i \ bounded m r (parent l (parent l x)) \ cmp (heap (parent l (parent l x))) (heap x))))" abbreviation is_heap_except_down :: "('a::order \ 'a::order \ bool) \ (int \ 'a::order) \ int \ int \ int \ bool" where "is_heap_except_down cmp heap l r i \ is_partial_heap_except_down cmp heap l l r i" text \As mentioned the notion of a partial heap is only needed for \make_heap\, which only uses \sift_down\ internally, so there doesn't need to be an additional definition for the partial heap version of the \sift_up\ invariant.\ definition is_heap_except_up :: "('a::order \ 'a::order \ bool) \ (int \ 'a::order) \ int \ int \ int \ bool" where "is_heap_except_up cmp heap l r i = (\ x. bounded l r x \ ((x \ i \ bounded l r (parent l x) \ cmp (heap (parent l x)) (heap x)) \ (parent l x = i \ bounded l r (parent l (parent l x)) \ cmp (heap (parent l (parent l x))) (heap x))))" subsubsection \Lemmas\ lemma empty_partial_heap[simp]: "is_partial_heap cmp heap l r r" unfolding is_partial_heap_def by linarith lemma is_partial_heap_smaller_back: "is_partial_heap cmp heap l m r \ r' \ r \ is_partial_heap cmp heap l m r'" unfolding is_partial_heap_def by simp lemma is_partial_heap_smaller_front: "is_partial_heap cmp heap l m r \ m \ m' \ is_partial_heap cmp heap l m' r" unfolding is_partial_heap_def by simp text \The second half of each array is a is a partial binary heap, since it contains only leafs, which are all trivial binary heaps.\ lemma snd_half_is_partial_heap: "(l + r) div 2 \ m \ is_partial_heap cmp heap l m r" unfolding is_partial_heap_def parent_def by linarith lemma modify_outside_partial_heap: assumes "heap = heap' on {m..The next few lemmas formalize how the heap invariant is weakened, when the heap is modified in a certain way.\ text \This lemma is used by \make_heap\.\ lemma partial_heap_added_first_el: assumes "l \ m" "m \ r" "is_partial_heap cmp heap l (m + 1) r" shows "is_partial_heap_except_down cmp heap l m r m" unfolding is_partial_heap_except_down_def proof fix x let ?p_x = "parent l x" let ?gp_x = "parent l ?p_x" show "bounded m r x \ (?p_x \ m \ bounded m r ?p_x \ cmp (heap ?p_x) (heap x)) \ (?p_x = m \ bounded m r ?gp_x \ cmp (heap ?gp_x) (heap x))" proof assume x_bound: "bounded m r x" have p_x_lower: "?p_x \ m \ bounded m r ?p_x \ ?p_x \ m + 1" by simp hence "?p_x \ m \ bounded m r ?p_x \ x \ m + 1" using parent_upper_bound[of l x] x_bound assms(1) by linarith hence p_invariant: "(?p_x \ m \ bounded m r ?p_x \ cmp (heap ?p_x) (heap x))" using assms(3) is_partial_heap_def p_x_lower x_bound by blast have gp_up_bound: "(?p_x = m \ ?gp_x < m)" by (simp add: assms(1) parent_upper_bound) show "(?p_x \ m \ bounded m r ?p_x \ cmp (heap ?p_x) (heap x)) \ (?p_x = m \ bounded m r ?gp_x \ cmp (heap ?gp_x) (heap x))" using gp_up_bound p_invariant by linarith qed qed text \This lemma is used by \del_min\.\ lemma heap_changed_first_el: assumes "is_heap cmp heap l r" "l \ r" shows "is_heap_except_down cmp (heap(l := b)) l r l" proof - have "is_partial_heap cmp heap l (l + 1) r" using assms(1) is_partial_heap_smaller_front by fastforce hence "is_partial_heap cmp (heap(l := b)) l (l + 1) r" using modify_outside_partial_heap[of heap] by simp thus ?thesis by (simp add: assms(2) partial_heap_added_first_el) qed text \This lemma is used by \insert\.\ lemma heap_appended_el: assumes "is_heap cmp heap l r" "heap = heap' on {l..First Heap Element\ text \The next step is to show that the first element of the heap is always the ``smallest'' according to the given comparison function. For the proof a rule for strong induction on lower bounded integers is needed. Its proof is based on the proof of strong induction on natural numbers found in \cite{Str_Ind}.\ lemma strong_int_gr_induct_helper: assumes "k < (i::int)" "(\i. k < i \ (\j. k < j \ j < i \ P j) \ P i)" shows "\j. k < j \ j < i \ P j" using assms proof(induction i rule: int_gr_induct) case base then show ?case by linarith next case (step i) then show ?case proof(cases "j = i") case True then show ?thesis using step.IH step.prems(1,3) by blast next case False hence "j < i" using step.prems(2) by simp then show ?thesis using step.IH step.prems(1,3) by blast qed qed theorem strong_int_gr_induct: assumes "k < (i::int)" "(\i. k < i \ (\j. k < j \ j < i \ P j) \ P i)" shows "P i" using assms less_induct strong_int_gr_induct_helper by blast text \Now the main theorem, that the first heap element is the ``smallest'' according to the given comparison function, can be proven.\ theorem heap_first_el: assumes "is_heap cmp heap l r" "transp cmp" "l < x" "x < r" shows "cmp (heap l) (heap x)" using assms unfolding is_partial_heap_def proof(induction x rule: strong_int_gr_induct[of l]) case 1 then show ?case using assms(3) by simp next case (2 i) have cmp_pi_i: "cmp (heap (parent l i)) (heap i)" using "2.hyps" "2.prems"(1,4) parent_bounds by simp then show ?case proof(cases) assume "parent l i > l" then have "cmp (heap l) (heap (parent l i))" using "2.IH" "2.prems"(1,2,4) parent_upper_bound_alt by simp then show ?thesis using "2.prems"(2) cmp_pi_i transpE by metis next assume "\ parent l i > l" then have "parent l i = l" using "2.hyps" dual_order.order_iff_strict parent_lower_bound by metis then show ?thesis using cmp_pi_i by simp qed qed section \General Lemmas on Arrays\ text \Some additional lemmas on @{const "mset_ran"}, @{const "swap"} and @{const "eq_on"} are needed for the final proofs.\ subsection \Lemmas on @{const "mset_ran"}\ abbreviation arr_mset :: "(int \ 'a) \ int \ int \ 'a multiset" where "arr_mset arr l r \ mset_ran arr {l..# (arr_mset arr l r) \ (\i. bounded l r i \ arr i = x)" unfolding mset_ran_def by fastforce lemma arr_mset_remove_last: "l \ r \ arr_mset arr l r = arr_mset arr l (r + 1) - {#arr r#}" by (simp add: intvs_upper_decr mset_ran_def) lemma arr_mset_append: "l \ r \ arr_mset arr l (r + 1) = arr_mset arr l r + {#arr r#}" using arr_mset_remove_last[of l r arr] by simp corollary arr_mset_append_alt: "l \ r \ arr_mset (arr(r := b)) l (r + 1) = arr_mset arr l r + {#b#}" by (simp add: arr_mset_append mset_ran_subst_outside) lemma arr_mset_remove_first: "i \ r \ arr_mset arr (i - 1) r = arr_mset arr i r + {#arr (i - 1)#}" by(induction r rule: int_ge_induct) (auto simp add: arr_mset_append) lemma arr_mset_split: assumes "l \ m" "m \ r" shows "arr_mset arr l r = arr_mset arr l m + arr_mset arr m r" using assms proof(induction m rule: int_ge_induct[of l]) case (step i) have add_last: "arr_mset arr l (i + 1) = arr_mset arr l i + {#arr i#}" using step arr_mset_append by blast have rem_first: "arr_mset arr (i+1) r = arr_mset arr i r - {#arr i#}" by (metis step.prems arr_mset_remove_first add_diff_cancel_right') show ?case using step add_last rem_first by fastforce qed (simp) text \That the first element in a heap is the ``smallest'', can now be expressed using multisets.\ corollary heap_first_el_alt: assumes "transp cmp" "is_heap cmp heap l r" "x \# (arr_mset heap l r)" "heap l \ x" shows "cmp (heap l) x" by(metis assms heap_first_el in_mset_imp_in_array le_less) subsection \Lemmas on @{term "swap"} and @{term "eq_on"}\ lemma eq_on_subset: "arr1 = arr2 on R \ S \ R \ arr1 = arr2 on S" by (simp add: eq_on_def set_mp) lemma swap_swaps: "arr' = swap arr x y \ arr' y = arr x \ arr' x = arr y" unfolding swap_def by simp lemma swap_only_swaps: "arr' = swap arr x y \ z \ x \ z \ y \ arr' z = arr z" unfolding swap_def by simp lemma swap_commute: "swap arr x y = swap arr y x" unfolding swap_def by fastforce lemma swap_eq_on: "arr1 = arr2 on S \ x \ S \ y \ S \ arr1 = swap arr2 x y on S" unfolding swap_def by simp corollary swap_parent_eq_on: assumes "arr1 = arr2 on - {l.. c = r_child l p" "l \ p" "c < r" shows "arr1 = swap arr2 p c on - {l.. c = r_child l p" "l \ p" "c < r" shows "arr_mset arr1 l r = arr_mset (swap arr2 p c) l r" proof - have child_bounded: "l < c \ c < r" by (metis assms(2-4) parent_lower_bound parent_two_children) have parent_bounded: "bounded l r p" by (metis assms(2-4) dual_order.strict_trans parent_two_children parent_upper_bound_alt) thus ?thesis using assms(1) child_bounded mset_ran_swap[of p "{l..The following lemma shows, which propositions have to hold on the pre-swap array, so that a comparison between two elements holds on the post-swap array. This is useful for the proofs of the loop invariants of \sift_up\ and \sift_down\. The lemma is kept quite general (except for the argument order) and could probably be more closely related to the parent relation for more concise proofs.\ lemma cmp_swapI: fixes arr::"'a::order \ 'a::order" assumes "m < n \ x < y" "m < n \ x < y \ x = m \ y = n \ P (arr n) (arr m)" "m < n \ x < y \ x \ m \ x \ n \ y \ m \ y \ n \ P (arr m) (arr n)" "m < n \ x < y \ x = m \ y \ m \ y \ n \ P (arr y) (arr n)" "m < n \ x < y \ x = n \ y \ m \ y \ n \ P (arr m) (arr y)" "m < n \ x < y \ x \ m \ x \ n \ y = n \ P (arr m) (arr x)" "m < n \ x < y \ x \ m \ x \ n \ y = m \ P (arr x) (arr n)" shows "P (swap arr x y m) (swap arr x y n)" by (metis assms order.asym swap_only_swaps swap_swaps) section \Imperative Heap Implementation\ text \The following imperative heap functions are based on \cite{MS} and \cite{CLRS}. All functions, that are recursive in these books, are iterative in the following implementations. The function definitions are done with IMP2 \cite{IMP2-AFP}. From now on the heaps only contain ints and only use \\\ as comparison function. The auxiliary lemmas used from now on are heavily modeled after the proof goals, that are generated by the vcg tool (also part of IMP2).\ subsection \Simple Functions\ subsubsection \Parent, Children and Swap\ text \In this section the parent and children relations are expressed as IMP2 procedures. Additionally a simple procedure, that swaps two array elements, is defined.\ procedure_spec prnt (l, x) returns p assumes True ensures "p = parent l\<^sub>0 x\<^sub>0" defines \p = ((x - l - 1) / 2 + l)\ by vcg (simp add: parent_def) procedure_spec left_child (l, x) returns lc assumes True ensures "lc = l_child l\<^sub>0 x\<^sub>0" defines \lc = 2 * x - l + 1\ by vcg (simp add: l_child_def) procedure_spec right_child (l, x) returns rc assumes True ensures "rc = r_child l\<^sub>0 x\<^sub>0" defines \rc = 2 * x - l + 2\ by vcg (simp add: r_child_def) procedure_spec swp (heap, x, y) returns heap assumes True ensures "heap = swap heap\<^sub>0 x\<^sub>0 y\<^sub>0 " defines \tmp = heap[x]; heap[x] = heap[y]; heap[y] = tmp\ by vcg (simp add: swap_def) subsubsection \\get_min\\ text \In this section \get_min\ is defined, which simply returns the first element (the minimum) of the heap. For this definition an additional theorem is proven, which enables the use of @{const "Min_mset"} in the postcondition.\ theorem heap_minimum: assumes "l < r" "is_heap (\) heap l r" shows "heap l = Min_mset (arr_mset heap l r)" proof - have "(\x \# (arr_mset heap l r). (heap l) \ x)" - using assms(2) heap_first_el_alt transp_le by blast + using assms(2) heap_first_el_alt transp_on_le by blast thus ?thesis by (simp add: assms(1) dual_order.antisym) qed procedure_spec get_min (heap, l, r) returns min assumes "l < r \ is_heap (\) heap l r" ensures "min = Min_mset (arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0)" for heap[] l r defines \min = heap[l]\ by vcg (simp add: heap_minimum) subsection \Modifying Functions\ subsubsection \\sift_up\ and \insert\\ text \The next heap function is \insert\, which internally uses \sift_up\. In the beginning of this section \sift_up_step\ is proven, which states that each \sift_up\ loop iteration correctly transforms the weakened heap invariant. For its proof two additional auxiliary lemmas are used. After \sift_up_step\ \sift_up\ and then \insert\ are verified.\ text \\sift_up_step\ can be proven directly by the smt-solver without auxiliary lemmas, but they were introduced to show the proof details. The analogous proofs for \sift_down\ were just solved with smt, since the proof structure should be very similar, even though the \sift_down\ proof goals are slightly more complex.\ lemma sift_up_step_aux1: fixes heap::"int \ int" assumes "is_heap_except_up (\) heap l r x" "parent l x \ l" "(heap x) \ (heap (parent l x))" "bounded l r k" "k \ (parent l x)" "bounded l r (parent l k)" shows "(swap heap (parent l x) x (parent l k)) \ (swap heap (parent l x) x k)" apply(intro cmp_swapI[of "(parent l k)" k "(parent l x)" x "(\)" heap]) subgoal using assms(2,6) parent_upper_bound_alt by blast subgoal using assms(3) by blast subgoal using assms(1,4,6) unfolding is_heap_except_up_def by blast subgoal using assms(1,3,4,6) unfolding is_heap_except_up_def by fastforce subgoal using assms(5) by blast subgoal by blast subgoal using assms(1,2,4) unfolding is_heap_except_up_def by simp done lemma sift_up_step_aux2: fixes heap::"int \ int" assumes "is_heap_except_up (\) heap l r x" "parent l x \ l" "heap x \ (heap (parent l x))" "bounded l r k" "parent l k = parent l x" "bounded l r (parent l (parent l k))" shows "swap heap (parent l x) x (parent l (parent l k)) \ swap heap (parent l x) x k" using assms unfolding is_heap_except_up_def proof- let ?gp_k = "parent l (parent l k)" let ?gp_x = "parent l (parent l x)" have gp_k_eq_gp_x: "swap heap (parent l x) x ?gp_k = heap ?gp_x" by (metis assms(2,5) grand_parent_upper_bound less_irrefl swap_only_swaps) show ?thesis using assms unfolding is_heap_except_up_def proof(cases) assume k_eq_x: "k = x" have "swap heap (parent l x) x k = heap (parent l x)" by (metis k_eq_x swap_swaps) then show ?thesis using assms(1,2,4,6) unfolding is_heap_except_up_def by (metis gp_k_eq_gp_x k_eq_x parent_bounds parent_lower_bound) next assume k_neq_x: "k \ x" have "swap heap (parent l x) x k = heap k" by (metis assms(5) gp_k_eq_gp_x k_neq_x swap_only_swaps) then show ?thesis using assms unfolding is_heap_except_up_def by (metis gp_k_eq_gp_x k_neq_x order_trans parent_bounds parent_lower_bound) qed qed lemma sift_up_step: fixes heap::"int \ int" assumes "is_heap_except_up (\) heap l r x" "parent l x \ l" "(heap x) \ (heap (parent l x))" shows "is_heap_except_up (\) (swap heap (parent l x) x) l r (parent l x)" using assms sift_up_step_aux1 sift_up_step_aux2 unfolding is_heap_except_up_def by blast text \\sift_up\ restores the heap invariant, that is only violated at the current position, by iteratively swapping the current element with its parent until the beginning of the array is reached or the current element is bigger than its parent.\ procedure_spec sift_up (heap, l, r, x) returns heap assumes "is_heap_except_up (\) heap l r x \ bounded l r x" ensures "is_heap (\) heap l\<^sub>0 r\<^sub>0 \ arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0 = arr_mset heap l\<^sub>0 r\<^sub>0 \ heap\<^sub>0 = heap on - {l\<^sub>0..0}" for heap[] l x r defines \ p = prnt(l, x); while (x > l \ heap[x] \ heap[p]) @variant \x - l\ @invariant \is_heap_except_up (\) heap l r x \ p = parent l x \ bounded l r x \ arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0 = arr_mset heap l r \ heap\<^sub>0 = heap on - {l.. { heap = swp(heap, p, x); x = p; p = prnt(l, x) }\ apply vcg_cs apply(intro conjI) subgoal using parent_lower_bound sift_up_step by blast subgoal using parent_lower_bound by blast subgoal using parent_bounds by blast subgoal using parent_bounds by (simp add: mset_ran_swap) subgoal using swap_parent_eq_on by blast subgoal using parent_upper_bound by simp subgoal unfolding is_heap_except_up_def is_partial_heap_def by (metis le_less not_less parent_lower_bound) done text \\insert\ inserts an element into a heap by appending it to the heap and restoring the heap invariant with @{const "sift_up"}.\ procedure_spec insert (heap, l, r, el) returns (heap, l, r) assumes "is_heap (\) heap l r \ l \ r" ensures "is_heap (\) heap l r \ arr_mset heap l r = arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0 + {#el\<^sub>0#} \ l = l\<^sub>0 \ r = r\<^sub>0 + 1 \ heap\<^sub>0 = heap on - {l.. heap[r] = el; x = r; r = r + 1; heap = sift_up(heap, l, r, x) \ apply vcg_cs subgoal by (simp add: heap_appended_el) subgoal by (metis arr_mset_append_alt add_mset_add_single) done subsubsection \\sift_down\, \del_min\ and \make_heap\\ text \The next heap functions are \del_min\ and \make_heap\, which both use \sift_down\ to restore/establish the heap invariant. \sift_down\ is proven first (this time without additional auxiliary lemmas) followed by \del_min\ and \make_heap\.\ text \\sift_down\ restores the heap invariant, that is only violated at the current position, by iteratively swapping the current element with its smallest child until the end of the array is reached or the current element is smaller than its children.\ procedure_spec sift_down(heap, l, r, x) returns heap assumes "is_partial_heap_except_down (\) heap l x r x \ l \ x \ x \ r" ensures "is_partial_heap (\) heap l\<^sub>0 x\<^sub>0 r\<^sub>0 \ arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0 = arr_mset heap l\<^sub>0 r\<^sub>0 \ heap\<^sub>0 = heap on - {l\<^sub>0..0}" defines \ lc = left_child(l, x); rc = right_child(l, x); while (lc < r \ (heap[lc] < heap[x] \ (rc < r \ heap[rc] < heap[x]))) @variant \r - x\ @invariant \is_partial_heap_except_down (\) heap l x\<^sub>0 r x \ x\<^sub>0 \ x \ x \ r \ lc = l_child l x \ rc = r_child l x \ arr_mset heap\<^sub>0 l r = arr_mset heap l r \ heap\<^sub>0 = heap on - {l.. { smallest = lc; if (rc < r \ heap[rc] < heap[lc]) { smallest = rc }; heap = swp(heap, x, smallest); x = smallest; lc = left_child(l, x); rc = right_child(l, x) }\ apply vcg_cs subgoal apply(intro conjI) subgoal unfolding is_partial_heap_except_down_def by (smt parent_two_children swap_swaps swap_only_swaps swap_commute parent_upper_bound_alt) subgoal using r_child_lower_bound_alt by fastforce subgoal using swap_child_mset order_trans by blast subgoal using swap_child_eq_on by fastforce done subgoal by (meson less_le_trans not_le order.asym r_child_lower_bound) subgoal apply(intro conjI) subgoal unfolding is_partial_heap_except_down_def by (smt parent_two_children swap_swaps swap_only_swaps swap_commute parent_upper_bound_alt) subgoal using l_child_lower_bound_alt by fastforce subgoal using swap_child_mset order_trans by blast subgoal using swap_child_eq_on by fastforce done subgoal by (meson less_le_trans not_le order.asym l_child_lower_bound) subgoal unfolding is_partial_heap_except_down_def is_partial_heap_def by (metis dual_order.strict_trans not_less parent_two_children smaller_l_child) done text \\del_min\ needs an additional lemma which shows, that it actually removes (only) the minimum from the heap.\ lemma del_min_mset: fixes heap::"int \ int" assumes "l < r" "is_heap (\) heap l r" "mod_heap = heap(l := heap (r - 1))" "arr_mset mod_heap l (r - 1) = arr_mset new_heap l (r - 1)" shows "arr_mset new_heap l (r - 1) = arr_mset heap l r - {#Min_mset (arr_mset heap l r)#}" proof - let ?heap_mset = "arr_mset heap l r" have l_is_min: "heap l = Min_mset ?heap_mset" using assms(1,2) heap_minimum by blast have "(arr_mset mod_heap l r) = ?heap_mset + {#heap (r-1)#} - {#heap l#}" by (simp add: assms(1,3) mset_ran_subst_inside) hence "(arr_mset mod_heap l (r - 1)) = ?heap_mset - {#heap l#}" by (simp add: assms(1,3) arr_mset_remove_last) thus ?thesis using assms(4) l_is_min by simp qed text \\del_min\ removes the minimum element from the heap by replacing the first element with the last element, shrinking the array by one and subsequently restoring the heap invariant with @{const "sift_down"}.\ procedure_spec del_min (heap, l, r) returns (heap, l, r) assumes "l < r \ is_heap (\) heap l r" ensures "is_heap (\) heap l r \ arr_mset heap l r = arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0 - {#Min_mset (arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0)#} \ l = l\<^sub>0 \ r = r\<^sub>0 - 1 \ heap\<^sub>0 = heap on - {l\<^sub>0..0}" for heap l r defines \ r = r - 1; heap[l] = heap[r]; heap = sift_down(heap, l, r, l) \ apply vcg_cs subgoal by (simp add: heap_changed_first_el is_partial_heap_smaller_back) subgoal apply(rule conjI) subgoal using del_min_mset by blast subgoal by (simp add: eq_on_def intvs_incdec(3) intvs_lower_incr) done done text \\make_heap\ transforms an arbitrary array into a heap by iterating through all array positions from the middle of the array up to the beginning of the array and calling @{const "sift_down"} for each one.\ procedure_spec make_heap (heap, l, r) returns heap assumes "l \ r" ensures "is_heap (\) heap l\<^sub>0 r\<^sub>0 \ arr_mset heap l\<^sub>0 r\<^sub>0 = arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0 \ heap\<^sub>0 = heap on - {l\<^sub>0..< r\<^sub>0}" for heap[] l r defines \ y = (r + l)/2 - 1; while (y \ l) @variant \y - l + 1\ @invariant \is_partial_heap (\) heap l (y + 1) r \ arr_mset heap l r = arr_mset heap\<^sub>0 l\<^sub>0 r\<^sub>0 \ l - 1 \ y \ y < r \ heap\<^sub>0 = heap on - {l.. { heap = sift_down(heap, l, r, y); y = y - 1 }\ apply(vcg_cs) subgoal apply(rule conjI) subgoal by (simp add: snd_half_is_partial_heap add.commute) subgoal by linarith done subgoal using partial_heap_added_first_el le_less by blast subgoal using eq_on_trans by blast subgoal using dual_order.antisym by fastforce done subsection \Heapsort Implementation\ text \The final part of this submission is the implementation of the in-place heapsort. Firstly it builds the \\\-heap and then it iteratively removes the minimum of the heap, which is put at the now vacant end of the shrinking heap. This is done until the heap is empty, which leaves the array sorted in descending order.\ subsubsection \Auxiliary Lemmas\ text \Firstly the notion of a sorted array is needed. This is more or less the same as @{const "ran_sorted"} generalized for arbitrary comparison functions.\ definition array_is_sorted :: "(int \ int \ bool) \ (int \ int) \ int \ int \ bool" where "array_is_sorted cmp a l r \ \i. \j. bounded l r i \ bounded l r j \ i < j \ cmp (a i) (a j)" text \This lemma states, that the heapsort doesn't change the elements contained in the array during the loop iterations.\ lemma heap_sort_mset_step: fixes arr::"int \ int" assumes "l < m" "m \ r" "arr_mset arr' l (m - 1) = arr_mset arr l m - {#Min_mset (arr_mset arr l m)#}" "arr = arr' on - {l..This lemma states, that each loop iteration leaves the growing second half of the array sorted in descending order.\ lemma heap_sort_second_half_sorted_step: fixes arr::"int \ int" assumes "l\<^sub>0 < m" "m \ r\<^sub>0" "arr = arr' on - {l\<^sub>0..i. \j. bounded m r\<^sub>0 i \ bounded m r\<^sub>0 j \ i < j \ arr j \ arr i" "\x\#arr_mset arr l\<^sub>0 m. \y\#arr_mset arr m r\<^sub>0. \ x < y" "bounded (m - 1) r\<^sub>0 i" "bounded (m - 1) r\<^sub>0 j" "i < j" "mod_arr = (arr'(m - 1 := Min_mset (arr_mset arr l\<^sub>0 m)))" shows "mod_arr j \ mod_arr i" proof - have second_half_eq: "mod_arr = arr on {m..< r\<^sub>0}" using assms(3, 9) unfolding eq_on_def by simp have j_stricter_bound: "bounded m r\<^sub>0 j" using assms(6-8) by simp then have el_at_j: "mod_arr j \# arr_mset arr m r\<^sub>0" using eq_onD second_half_eq by fastforce then show ?thesis proof(cases) assume "i = (m-1)" then have "mod_arr i \# arr_mset arr l\<^sub>0 m" by (simp add: assms(1, 9)) then show ?thesis using assms(5) el_at_j not_less by blast next assume "i \ (m-1)" then have "bounded m r\<^sub>0 i" using assms(6) by simp then show ?thesis using assms(4, 8) eq_on_def j_stricter_bound second_half_eq by force qed qed text \The following lemma shows that all elements in the first part of the array (the binary heap) are bigger than the elements in the second part (the sorted part) after every iteration. This lemma and the invariant of the \heap_sort\ loop use \\ x < y\ instead of \x \ y\ since \vcg_cs\ doesn't terminate in the latter case.\ lemma heap_sort_fst_part_bigger_snd_part_step: fixes arr::"int \ int" assumes "l\<^sub>0 < m" "m \ r\<^sub>0" "arr_mset arr' l\<^sub>0 (m - 1) = arr_mset arr l\<^sub>0 m - {#Min_mset (arr_mset arr l\<^sub>0 m)#}" "arr = arr' on - {l\<^sub>0..x\#arr_mset arr l\<^sub>0 m. \y\#arr_mset arr m r\<^sub>0. \ x < y" "mod_arr = arr'(m - 1 := Min_mset (arr_mset arr l\<^sub>0 m))" "x\#arr_mset mod_arr l\<^sub>0 (m - 1)" "y\#arr_mset mod_arr (m - 1) r\<^sub>0" shows "\ x < y" proof - have "{m..0} \ - {l\<^sub>0..0}" using assms(4) eq_on_sym eq_on_subset by blast hence arr_eq_on: "mod_arr = arr on {m..0}" by (simp add: assms(6)) hence same_mset: "arr_mset mod_arr m r\<^sub>0 = arr_mset arr m r\<^sub>0" using mset_ran_cong by blast have "x \# arr_mset arr l\<^sub>0 m" using same_mset assms by (metis assms(3,6,7) add_mset_remove_trivial_eq lran_upd_outside(2) mset_lran cancel_ab_semigroup_add_class.diff_right_commute diff_single_trivial multi_self_add_other_not_self order_refl) then have x_bigger_min: "x \ Min_mset (arr_mset arr l\<^sub>0 m)" using Min_le by blast have y_smaller_min: "y \ Min_mset (arr_mset arr l\<^sub>0 m)" proof(cases "y = mod_arr (m - 1)") case False hence "y\#arr_mset mod_arr (m - 1) r\<^sub>0 - {#mod_arr (m - 1)#}" by (metis assms(8) diff_single_trivial insert_DiffM insert_noteq_member) then have "y\#arr_mset arr m r\<^sub>0" by (simp add: assms(2) intvs_decr_l mset_ran_insert same_mset) then show ?thesis using assms(1) assms(5) by fastforce qed (simp add: assms(6)) then show ?thesis using x_bigger_min by linarith qed subsubsection \Implementation\ text \Now finally the correctness of the \heap_sort\ is shown. As mentioned, it starts by transforming the array into a minimum heap using @{const "make_heap"}. Then in each iteration it removes the first element from the heap with @{const "del_min"} after its value was retrieved with @{const "get_min"}. This value is then put at the position freed by @{const "del_min"}.\ program_spec heap_sort assumes "l \ r" ensures "array_is_sorted (\) arr l\<^sub>0 r\<^sub>0 \ arr_mset arr\<^sub>0 l\<^sub>0 r\<^sub>0 = arr_mset arr l\<^sub>0 r\<^sub>0 \ arr\<^sub>0 = arr on - {l\<^sub>0 ..0 } \ l = l\<^sub>0 \ r = r\<^sub>0" for l r arr[] defines \ arr = make_heap(arr, l, r); m = r; while (m > l) @variant \m - l + 1\ @invariant \is_heap (\) arr l m \ array_is_sorted (\) arr m r\<^sub>0 \ (\x \# arr_mset arr l\<^sub>0 m. \y \# arr_mset arr m r\<^sub>0. \ x < y) \ arr_mset arr\<^sub>0 l\<^sub>0 r\<^sub>0 = arr_mset arr l\<^sub>0 r\<^sub>0 \ l \ m \ m \ r\<^sub>0 \ l = l\<^sub>0 \ arr\<^sub>0 = arr on - {l\<^sub>0 ..0}\ { min = get_min(arr, l, m); (arr, l, m) = del_min(arr, l, m); arr[m] = min } \ apply vcg_cs subgoal unfolding array_is_sorted_def by simp subgoal apply(intro conjI) subgoal unfolding is_partial_heap_def by simp subgoal unfolding array_is_sorted_def using heap_sort_second_half_sorted_step by blast subgoal using heap_sort_fst_part_bigger_snd_part_step by blast subgoal using heap_sort_mset_step by blast subgoal unfolding eq_on_def by (metis ComplD ComplI atLeastLessThan_iff less_le_trans) done done end diff --git a/thys/Polynomials/MPoly_Type_Class_OAlist.thy b/thys/Polynomials/MPoly_Type_Class_OAlist.thy --- a/thys/Polynomials/MPoly_Type_Class_OAlist.thy +++ b/thys/Polynomials/MPoly_Type_Class_OAlist.thy @@ -1,963 +1,963 @@ (* Author: Fabian Immler, TU Muenchen *) (* Author: Florian Haftmann, TU Muenchen *) (* Author: Andreas Lochbihler, ETH Zurich *) (* Author: Alexander Maletzky, RISC Linz *) section \Executable Representation of Polynomial Mappings as Association Lists\ theory MPoly_Type_Class_OAlist imports Term_Order begin instantiation pp :: (type, "{equal, zero}") equal begin definition equal_pp :: "('a, 'b) pp \ ('a, 'b) pp \ bool" where "equal_pp p q \ (\t. lookup_pp p t = lookup_pp q t)" instance by standard (auto simp: equal_pp_def intro: pp_eqI) end instantiation poly_mapping :: (type, "{equal, zero}") equal begin definition equal_poly_mapping :: "('a, 'b) poly_mapping \ ('a, 'b) poly_mapping \ bool" where equal_poly_mapping_def [code del]: "equal_poly_mapping p q \ (\t. lookup p t = lookup q t)" instance by standard (auto simp: equal_poly_mapping_def intro: poly_mapping_eqI) end subsection \Power-Products Represented by @{type oalist_tc}\ definition PP_oalist :: "('a::linorder, 'b::zero) oalist_tc \ ('a, 'b) pp" where "PP_oalist xs = pp_of_fun (OAlist_tc_lookup xs)" code_datatype PP_oalist lemma lookup_PP_oalist [simp, code]: "lookup_pp (PP_oalist xs) = OAlist_tc_lookup xs" unfolding PP_oalist_def proof (rule lookup_pp_of_fun) have "{x. OAlist_tc_lookup xs x \ 0} \ fst ` set (list_of_oalist_tc xs)" proof (rule, simp) fix x assume "OAlist_tc_lookup xs x \ 0" thus "x \ fst ` set (list_of_oalist_tc xs)" using in_OAlist_tc_sorted_domain_iff_lookup set_OAlist_tc_sorted_domain by blast qed also have "finite ..." by simp finally (finite_subset) show "finite {x. OAlist_tc_lookup xs x \ 0}" . qed lemma keys_PP_oalist [code]: "keys_pp (PP_oalist xs) = set (OAlist_tc_sorted_domain xs)" by (rule set_eqI, simp add: keys_pp_iff in_OAlist_tc_sorted_domain_iff_lookup) lemma lex_comp_PP_oalist [code]: "lex_comp' (PP_oalist xs) (PP_oalist ys) = the (OAlist_tc_lex_ord (\_ x y. Some (comparator_of x y)) xs ys)" for xs ys::"('a::nat, 'b::nat) oalist_tc" proof (cases "lex_comp' (PP_oalist xs) (PP_oalist ys) = Eq") case True hence "PP_oalist xs = PP_oalist ys" by (rule lex_comp'_EqD) hence eq: "OAlist_tc_lookup xs = OAlist_tc_lookup ys" by (simp add: pp_eq_iff) have "OAlist_tc_lex_ord (\_ x y. Some (comparator_of x y)) xs ys = Some Eq" by (rule OAlist_tc_lex_ord_EqI, simp add: eq) thus ?thesis by (simp add: True) next case False then obtain x where 1: "x \ keys_pp (rep_nat_pp (PP_oalist xs)) \ keys_pp (rep_nat_pp (PP_oalist ys))" and 2: "comparator_of (lookup_pp (rep_nat_pp (PP_oalist xs)) x) (lookup_pp (rep_nat_pp (PP_oalist ys)) x) = lex_comp' (PP_oalist xs) (PP_oalist ys)" and 3: "\y. y < x \ lookup_pp (rep_nat_pp (PP_oalist xs)) y = lookup_pp (rep_nat_pp (PP_oalist ys)) y" by (rule lex_comp'_valE, blast) have "OAlist_tc_lex_ord (\_ x y. Some (comparator_of x y)) xs ys = Some (lex_comp' (PP_oalist xs) (PP_oalist ys))" proof (rule OAlist_tc_lex_ord_valI) from False show "Some (lex_comp' (PP_oalist xs) (PP_oalist ys)) \ Some Eq" by simp next from 1 have "abs_nat x \ abs_nat ` (keys_pp (rep_nat_pp (PP_oalist xs)) \ keys_pp (rep_nat_pp (PP_oalist ys)))" by (rule imageI) also have "... = fst ` set (list_of_oalist_tc xs) \ fst ` set (list_of_oalist_tc ys)" by (simp add: keys_rep_nat_pp_pp keys_PP_oalist OAlist_tc_sorted_domain_def image_Un image_image) finally show "abs_nat x \ fst ` set (list_of_oalist_tc xs) \ fst ` set (list_of_oalist_tc ys)" . next show "Some (lex_comp' (PP_oalist xs) (PP_oalist ys)) = Some (comparator_of (OAlist_tc_lookup xs (abs_nat x)) (OAlist_tc_lookup ys (abs_nat x)))" by (simp add: 2[symmetric] lookup_rep_nat_pp_pp) next fix y::'a assume "y < abs_nat x" hence "rep_nat y < x" by (metis abs_inverse ord_iff(2)) hence "lookup_pp (rep_nat_pp (PP_oalist xs)) (rep_nat y) = lookup_pp (rep_nat_pp (PP_oalist ys)) (rep_nat y)" by (rule 3) hence "OAlist_tc_lookup xs y = OAlist_tc_lookup ys y" by (auto simp: lookup_rep_nat_pp_pp elim: rep_inj) thus "Some (comparator_of (OAlist_tc_lookup xs y) (OAlist_tc_lookup ys y)) = Some Eq" by simp qed thus ?thesis by simp qed lemma zero_PP_oalist [code]: "(0::('a::linorder, 'b::zero) pp) = PP_oalist OAlist_tc_empty" by (rule pp_eqI, simp add: lookup_OAlist_tc_empty) lemma plus_PP_oalist [code]: "PP_oalist xs + PP_oalist ys = PP_oalist (OAlist_tc_map2_val_neutr (\_. (+)) xs ys)" by (rule pp_eqI, simp add: lookup_plus_pp, rule lookup_OAlist_tc_map2_val_neutr[symmetric], simp_all) lemma minus_PP_oalist [code]: "PP_oalist xs - PP_oalist ys = PP_oalist (OAlist_tc_map2_val_rneutr (\_. (-)) xs ys)" by (rule pp_eqI, simp add: lookup_minus_pp, rule lookup_OAlist_tc_map2_val_rneutr[symmetric], simp) lemma equal_PP_oalist [code]: "equal_class.equal (PP_oalist xs) (PP_oalist ys) = (xs = ys)" by (simp add: equal_eq pp_eq_iff, auto elim: OAlist_tc_lookup_inj) lemma lcs_PP_oalist [code]: "lcs (PP_oalist xs) (PP_oalist ys) = PP_oalist (OAlist_tc_map2_val_neutr (\_. max) xs ys)" for xs ys :: "('a::linorder, 'b::add_linorder_min) oalist_tc" by (rule pp_eqI, simp add: lookup_lcs_pp, rule lookup_OAlist_tc_map2_val_neutr[symmetric], simp_all add: max_def) lemma deg_pp_PP_oalist [code]: "deg_pp (PP_oalist xs) = sum_list (map snd (list_of_oalist_tc xs))" proof - have "irreflp ((<)::_::linorder \ _)" by (rule irreflpI, simp) have "deg_pp (PP_oalist xs) = sum (OAlist_tc_lookup xs) (set (OAlist_tc_sorted_domain xs))" by (simp add: deg_pp_alt keys_PP_oalist) also have "... = sum_list (map (OAlist_tc_lookup xs) (OAlist_tc_sorted_domain xs))" by (rule sum.distinct_set_conv_list, rule distinct_sorted_wrt_irrefl, - fact, fact transp_less, fact sorted_OAlist_tc_sorted_domain) + fact, fact transp_on_less, fact sorted_OAlist_tc_sorted_domain) also have "... = sum_list (map snd (list_of_oalist_tc xs))" by (rule arg_cong[where f=sum_list], simp add: OAlist_tc_sorted_domain_def OAlist_tc_lookup_eq_valueI) finally show ?thesis . qed lemma single_PP_oalist [code]: "single_pp x e = PP_oalist (oalist_tc_of_list [(x, e)])" by (rule pp_eqI, simp add: lookup_single_pp OAlist_tc_lookup_single) definition adds_pp_add_linorder :: "('b, 'a::add_linorder) pp \ _ \ bool" where [code_abbrev]: "adds_pp_add_linorder = (adds)" lemma adds_pp_PP_oalist [code]: "adds_pp_add_linorder (PP_oalist xs) (PP_oalist ys) = OAlist_tc_prod_ord (\_. less_eq) xs ys" for xs ys::"('a::linorder, 'b::add_linorder_min) oalist_tc" proof (simp add: adds_pp_add_linorder_def adds_pp_iff adds_poly_mapping lookup_pp.rep_eq[symmetric] OAlist_tc_prod_ord_alt le_fun_def, intro iffI allI ballI) fix k assume "\x. OAlist_tc_lookup xs x \ OAlist_tc_lookup ys x" thus "OAlist_tc_lookup xs k \ OAlist_tc_lookup ys k" by blast next fix x assume *: "\k\fst ` set (list_of_oalist_tc xs) \ fst ` set (list_of_oalist_tc ys). OAlist_tc_lookup xs k \ OAlist_tc_lookup ys k" show "OAlist_tc_lookup xs x \ OAlist_tc_lookup ys x" proof (cases "x \ fst ` set (list_of_oalist_tc xs) \ fst ` set (list_of_oalist_tc ys)") case True with * show ?thesis .. next case False hence "x \ set (OAlist_tc_sorted_domain xs)" and "x \ set (OAlist_tc_sorted_domain ys)" by (simp_all add: set_OAlist_tc_sorted_domain) thus ?thesis by (simp add: in_OAlist_tc_sorted_domain_iff_lookup) qed qed subsubsection \Constructor\ definition "sparse\<^sub>0 xs = PP_oalist (oalist_tc_of_list xs)" \\sparse representation\ subsubsection \Computations\ experiment begin abbreviation "X \ 0::nat" abbreviation "Y \ 1::nat" abbreviation "Z \ 2::nat" value [code] "sparse\<^sub>0 [(X, 2::nat), (Z, 7)]" lemma "sparse\<^sub>0 [(X, 2::nat), (Z, 7)] - sparse\<^sub>0 [(X, 2), (Z, 2)] = sparse\<^sub>0 [(Z, 5)]" by eval lemma "lcs (sparse\<^sub>0 [(X, 2::nat), (Y, 1), (Z, 7)]) (sparse\<^sub>0 [(Y, 3), (Z, 2)]) = sparse\<^sub>0 [(X, 2), (Y, 3), (Z, 7)]" by eval lemma "(sparse\<^sub>0 [(X, 2::nat), (Z, 1)]) adds (sparse\<^sub>0 [(X, 3), (Y, 2), (Z, 1)])" by eval lemma "lookup_pp (sparse\<^sub>0 [(X, 2::nat), (Z, 3)]) X = 2" by eval lemma "deg_pp (sparse\<^sub>0 [(X, 2::nat), (Y, 1), (Z, 3), (X, 1)]) = 6" by eval lemma "lex_comp (sparse\<^sub>0 [(X, 2::nat), (Y, 1), (Z, 3)]) (sparse\<^sub>0 [(X, 4)]) = Lt" by eval lemma "lex_comp (sparse\<^sub>0 [(X, 2::nat), (Y, 1), (Z, 3)], 3::nat) (sparse\<^sub>0 [(X, 4)], 2) = Lt" by eval lemma "lex_pp (sparse\<^sub>0 [(X, 2::nat), (Y, 1), (Z, 3)]) (sparse\<^sub>0 [(X, 4)])" by eval lemma "lex_pp (sparse\<^sub>0 [(X, 2::nat), (Y, 1), (Z, 3)]) (sparse\<^sub>0 [(X, 4)])" by eval lemma "\ dlex_pp (sparse\<^sub>0 [(X, 2::nat), (Y, 1), (Z, 3)]) (sparse\<^sub>0 [(X, 4)])" by eval lemma "dlex_pp (sparse\<^sub>0 [(X, 2::nat), (Y, 1), (Z, 2)]) (sparse\<^sub>0 [(X, 5)])" by eval lemma "\ drlex_pp (sparse\<^sub>0 [(X, 2::nat), (Y, 1), (Z, 2)]) (sparse\<^sub>0 [(X, 5)])" by eval end subsection \\MP_oalist\\ lift_definition MP_oalist :: "('a::nat_term, 'b::zero) oalist_ntm \ 'a \\<^sub>0 'b" is OAlist_lookup_ntm proof - fix xs :: "('a, 'b) oalist_ntm" have "{x. OAlist_lookup_ntm xs x \ 0} \ fst ` set (fst (list_of_oalist_ntm xs))" proof (rule, simp) fix x assume "OAlist_lookup_ntm xs x \ 0" thus "x \ fst ` set (fst (list_of_oalist_ntm xs))" using oa_ntm.in_sorted_domain_iff_lookup oa_ntm.set_sorted_domain by blast qed also have "finite ..." by simp finally (finite_subset) show "finite {x. OAlist_lookup_ntm xs x \ 0}" . qed lemmas [simp, code] = MP_oalist.rep_eq code_datatype MP_oalist lemma keys_MP_oalist [code]: "keys (MP_oalist xs) = set (map fst (fst (list_of_oalist_ntm xs)))" by (rule set_eqI, simp add: in_keys_iff oa_ntm.in_sorted_domain_iff_lookup[simplified oa_ntm.set_sorted_domain]) lemma MP_oalist_empty [simp]: "MP_oalist (OAlist_empty_ntm ko) = 0" by (rule poly_mapping_eqI, simp add: oa_ntm.lookup_empty) lemma zero_MP_oalist [code]: "(0::('a::{linorder,nat_term} \\<^sub>0 'b::zero)) = MP_oalist (OAlist_empty_ntm nat_term_order_of_le)" by simp definition is_zero :: "('a \\<^sub>0 'b::zero) \ bool" where [code_abbrev]: "is_zero p \ (p = 0)" lemma is_zero_MP_oalist [code]: "is_zero (MP_oalist xs) = List.null (fst (list_of_oalist_ntm xs))" unfolding is_zero_def List.null_def proof assume "MP_oalist xs = 0" hence "OAlist_lookup_ntm xs k = 0" for k by (simp add: poly_mapping_eq_iff) thus "fst (list_of_oalist_ntm xs) = []" by (metis image_eqI ko_ntm.min_key_val_raw_in oa_ntm.in_sorted_domain_iff_lookup oa_ntm.set_sorted_domain) next assume "fst (list_of_oalist_ntm xs) = []" hence "OAlist_lookup_ntm xs k = 0" for k by (metis oa_ntm.list_of_oalist_empty oa_ntm.lookup_empty oalist_ntm_eqI surjective_pairing) thus "MP_oalist xs = 0" by (simp add: poly_mapping_eq_iff ext) qed lemma plus_MP_oalist [code]: "MP_oalist xs + MP_oalist ys = MP_oalist (OAlist_map2_val_neutr_ntm (\_. (+)) xs ys)" by (rule poly_mapping_eqI, simp add: lookup_plus_fun, rule oa_ntm.lookup_map2_val_neutr[symmetric], simp_all) lemma minus_MP_oalist [code]: "MP_oalist xs - MP_oalist ys = MP_oalist (OAlist_map2_val_rneutr_ntm (\_. (-)) xs ys)" by (rule poly_mapping_eqI, simp add: lookup_minus_fun, rule oa_ntm.lookup_map2_val_rneutr[symmetric], simp) lemma uminus_MP_oalist [code]: "- MP_oalist xs = MP_oalist (OAlist_map_val_ntm (\_. uminus) xs)" by (rule poly_mapping_eqI, simp, rule oa_ntm.lookup_map_val[symmetric], simp) lemma equal_MP_oalist [code]: "equal_class.equal (MP_oalist xs) (MP_oalist ys) = (OAlist_eq_ntm xs ys)" by (simp add: oa_ntm.oalist_eq_alt equal_eq poly_mapping_eq_iff) lemma map_MP_oalist [code]: "Poly_Mapping.map f (MP_oalist xs) = MP_oalist (OAlist_map_val_ntm (\_. f) xs)" proof - have eq: "OAlist_map_val_ntm (\_. f) xs = OAlist_map_val_ntm (\_ c. f c when c \ 0) xs" proof (rule oa_ntm.map_val_cong) fix t c assume *: "(t, c) \ set (fst (list_of_oalist_ntm xs))" hence "fst (t, c) \ fst ` set (fst (list_of_oalist_ntm xs))" by (rule imageI) hence "OAlist_lookup_ntm xs t \ 0" by (simp add: oa_ntm.in_sorted_domain_iff_lookup[simplified oa_ntm.set_sorted_domain]) moreover from * have "OAlist_lookup_ntm xs t = c" by (rule oa_ntm.lookup_eq_valueI) ultimately have "c \ 0" by simp thus "f c = (f c when c \ 0)" by simp qed show ?thesis by (rule poly_mapping_eqI, simp add: Poly_Mapping.map.rep_eq eq, rule oa_ntm.lookup_map_val[symmetric], simp) qed lemma range_MP_oalist [code]: "Poly_Mapping.range (MP_oalist xs) = set (map snd (fst (list_of_oalist_ntm xs)))" proof (simp add: Poly_Mapping.range.rep_eq, intro set_eqI iffI) fix c assume "c \ range (OAlist_lookup_ntm xs) - {0}" hence "c \ range (OAlist_lookup_ntm xs)" and "c \ 0" by simp_all from this(1) obtain t where "OAlist_lookup_ntm xs t = c" by fastforce with \c \ 0\ have "(t, c) \ set (fst (list_of_oalist_ntm xs))" by (simp add: oa_ntm.lookup_eq_value) hence "snd (t, c) \ snd ` set (fst (list_of_oalist_ntm xs))" by (rule imageI) thus "c \ snd ` set (fst (list_of_oalist_ntm xs))" by simp next fix c assume "c \ snd ` set (fst (list_of_oalist_ntm xs))" then obtain t where *: "(t, c) \ set (fst (list_of_oalist_ntm xs))" by fastforce hence "fst (t, c) \ fst ` set (fst (list_of_oalist_ntm xs))" by (rule imageI) hence "OAlist_lookup_ntm xs t \ 0" by (simp add: oa_ntm.in_sorted_domain_iff_lookup[simplified oa_ntm.set_sorted_domain]) moreover from * have "OAlist_lookup_ntm xs t = c" by (rule oa_ntm.lookup_eq_valueI) ultimately show "c \ range (OAlist_lookup_ntm xs) - {0}" by fastforce qed lemma if_poly_mapping_eq_iff: "(if x = y then a else b) = (if (\i\keys x \ keys y. lookup x i = lookup y i) then a else b)" by simp (metis UnI1 UnI2 in_keys_iff poly_mapping_eqI) lemma keys_add_eq: "keys (a + b) = keys a \ keys b - {x \ keys a \ keys b. lookup a x + lookup b x = 0}" by (auto simp: in_keys_iff lookup_add add_eq_0_iff simp del: lookup_not_eq_zero_eq_in_keys) locale gd_nat_term = gd_term pair_of_term term_of_pair "\s t. le_of_nat_term_order cmp_term (term_of_pair (s, the_min)) (term_of_pair (t, the_min))" "\s t. lt_of_nat_term_order cmp_term (term_of_pair (s, the_min)) (term_of_pair (t, the_min))" "le_of_nat_term_order cmp_term" "lt_of_nat_term_order cmp_term" for pair_of_term::"'t::nat_term \ ('a::{nat_term,graded_dickson_powerprod} \ 'k::{countable,the_min,wellorder})" and term_of_pair::"('a \ 'k) \ 't" and cmp_term + assumes splus_eq_splus: "t \ u = nat_term_class.splus (term_of_pair (t, the_min)) u" begin definition shift_map_keys :: "'a \ ('b \ 'b) \ ('t, 'b) oalist_ntm \ ('t, 'b::semiring_0) oalist_ntm" where "shift_map_keys t f xs = OAlist_ntm (map_raw (\kv. (t \ fst kv, f (snd kv))) (list_of_oalist_ntm xs))" lemma list_of_oalist_shift_keys: "list_of_oalist_ntm (shift_map_keys t f xs) = (map_raw (\kv. (t \ fst kv, f (snd kv))) (list_of_oalist_ntm xs))" unfolding shift_map_keys_def by (rule oa_ntm.list_of_oalist_of_list_id, rule ko_ntm.oalist_inv_map_raw, fact oalist_inv_list_of_oalist_ntm, simp add: nat_term_compare_inv_conv[symmetric] nat_term_compare_inv_def splus_eq_splus nat_term_compare_splus) lemma lookup_shift_map_keys_plus: "lookup (MP_oalist (shift_map_keys t ((*) c) xs)) (t \ u) = c * lookup (MP_oalist xs) u" (is "?l = ?r") proof - let ?f = "\kv. (t \ fst kv, c * snd kv)" have "?l = lookup_ko_ntm (map_raw ?f (list_of_oalist_ntm xs)) (fst (?f (u, c)))" by (simp add: oa_ntm.lookup_def list_of_oalist_shift_keys) also have "... = snd (?f (u, lookup_ko_ntm (list_of_oalist_ntm xs) u))" by (rule ko_ntm.lookup_raw_map_raw, fact oalist_inv_list_of_oalist_ntm, simp, simp add: nat_term_compare_inv_conv[symmetric] nat_term_compare_inv_def splus_eq_splus nat_term_compare_splus) also have "... = ?r" by (simp add: oa_ntm.lookup_def) finally show ?thesis . qed lemma keys_shift_map_keys_subset: "keys (MP_oalist (shift_map_keys t ((*) c) xs)) \ ((\) t) ` keys (MP_oalist xs)" (is "?l \ ?r") proof - let ?f = "\kv. (t \ fst kv, c * snd kv)" have "?l = fst ` set (fst (map_raw ?f (list_of_oalist_ntm xs)))" by (simp add: keys_MP_oalist list_of_oalist_shift_keys) also from ko_ntm.map_raw_subset have "... \ fst ` ?f ` set (fst (list_of_oalist_ntm xs))" by (rule image_mono) also have "... \ ?r" by (simp add: keys_MP_oalist image_image) finally show ?thesis . qed lemma monom_mult_MP_oalist [code]: "monom_mult c t (MP_oalist xs) = MP_oalist (if c = 0 then OAlist_empty_ntm (snd (list_of_oalist_ntm xs)) else shift_map_keys t ((*) c) xs)" proof (cases "c = 0") case True hence "monom_mult c t (MP_oalist xs) = 0" using monom_mult_zero_left by simp thus ?thesis using True by simp next case False have "monom_mult c t (MP_oalist xs) = MP_oalist (shift_map_keys t ((*) c) xs)" proof (rule poly_mapping_eqI, simp add: lookup_monom_mult del: MP_oalist.rep_eq, intro conjI impI) fix u assume "t adds\<^sub>p u" then obtain v where "u = t \ v" by (rule adds_ppE) thus "c * lookup (MP_oalist xs) (u \ t) = lookup (MP_oalist (shift_map_keys t ((*) c) xs)) u" by (simp add: splus_sminus lookup_shift_map_keys_plus del: MP_oalist.rep_eq) next fix u assume "\ t adds\<^sub>p u" have "u \ keys (MP_oalist (shift_map_keys t ((*) c) xs))" proof assume "u \ keys (MP_oalist (shift_map_keys t ((*) c) xs))" also have "... \ ((\) t) ` keys (MP_oalist xs)" by (fact keys_shift_map_keys_subset) finally obtain v where "u = t \ v" .. hence "t adds\<^sub>p u" by (rule adds_ppI) with \\ t adds\<^sub>p u\ show False .. qed thus "lookup (MP_oalist (shift_map_keys t ((*) c) xs)) u = 0" by (simp add: in_keys_iff) qed thus ?thesis by (simp add: False) qed lemma mult_scalar_MP_oalist [code]: "(MP_oalist xs) \ (MP_oalist ys) = (if is_zero (MP_oalist xs) then MP_oalist (OAlist_empty_ntm (snd (list_of_oalist_ntm ys))) else let ct = OAlist_hd_ntm xs in monom_mult (snd ct) (fst ct) (MP_oalist ys) + (MP_oalist (OAlist_tl_ntm xs)) \ (MP_oalist ys))" proof (split if_split, intro conjI impI) assume "is_zero (MP_oalist xs)" thus "MP_oalist xs \ MP_oalist ys = MP_oalist (OAlist_empty_ntm (snd (list_of_oalist_ntm ys)))" by (simp add: is_zero_def) next assume "\ is_zero (MP_oalist xs)" hence *: "fst (list_of_oalist_ntm xs) \ []" by (simp add: is_zero_MP_oalist List.null_def) define ct where "ct = OAlist_hd_ntm xs" have eq: "except (MP_oalist xs) {fst ct} = MP_oalist (OAlist_tl_ntm xs)" by (rule poly_mapping_eqI, simp add: lookup_except ct_def oa_ntm.lookup_tl') have "MP_oalist xs \ MP_oalist ys = monom_mult (lookup (MP_oalist xs) (fst ct)) (fst ct) (MP_oalist ys) + except (MP_oalist xs) {fst ct} \ MP_oalist ys" by (fact mult_scalar_rec_left) also have "... = monom_mult (snd ct) (fst ct) (MP_oalist ys) + except (MP_oalist xs) {fst ct} \ MP_oalist ys" using * by (simp add: ct_def oa_ntm.snd_hd) also have "... = monom_mult (snd ct) (fst ct) (MP_oalist ys) + MP_oalist (OAlist_tl_ntm xs) \ MP_oalist ys" by (simp only: eq) finally show "MP_oalist xs \ MP_oalist ys = (let ct = OAlist_hd_ntm xs in monom_mult (snd ct) (fst ct) (MP_oalist ys) + MP_oalist (OAlist_tl_ntm xs) \ MP_oalist ys)" by (simp add: ct_def Let_def) qed end (* ordered_nat_term *) subsubsection \Special case of addition: adding monomials\ definition plus_monomial_less :: "('a \\<^sub>0 'b) \ 'b \ 'a \ ('a \\<^sub>0 'b::monoid_add)" where "plus_monomial_less p c u = p + monomial c u" text \@{const plus_monomial_less} is useful when adding a monomial to a polynomial, where the term of the monomial is known to be smaller than all terms in the polynomial, because it can be implemented more efficiently than general addition.\ lemma plus_monomial_less_MP_oalist [code]: "plus_monomial_less (MP_oalist xs) c u = MP_oalist (OAlist_update_by_fun_gr_ntm u (\c0. c0 + c) xs)" unfolding plus_monomial_less_def oa_ntm.update_by_fun_gr_eq_update_by_fun by (rule poly_mapping_eqI, simp add: lookup_plus_fun oa_ntm.lookup_update_by_fun lookup_single) text \@{const plus_monomial_less} is computed by @{const OAlist_update_by_fun_gr_ntm}, because greater terms come @{emph \before\} smaller ones in @{type oalist_ntm}.\ subsubsection \Constructors\ definition "distr\<^sub>0 ko xs = MP_oalist (oalist_of_list_ntm (xs, ko))" \\sparse representation\ definition V\<^sub>0 :: "'a \ ('a, nat) pp \\<^sub>0 'b::{one,zero}" where "V\<^sub>0 n \ monomial 1 (single_pp n 1)" definition C\<^sub>0 :: "'b \ ('a, nat) pp \\<^sub>0 'b::zero" where "C\<^sub>0 c \ monomial c 0" lemma C\<^sub>0_one: "C\<^sub>0 1 = 1" by (simp add: C\<^sub>0_def) lemma C\<^sub>0_numeral: "C\<^sub>0 (numeral x) = numeral x" by (auto intro!: poly_mapping_eqI simp: C\<^sub>0_def lookup_numeral) lemma C\<^sub>0_minus: "C\<^sub>0 (- x) = - C\<^sub>0 x" by (simp add: C\<^sub>0_def single_uminus) lemma C\<^sub>0_zero: "C\<^sub>0 0 = 0" by (auto intro!: poly_mapping_eqI simp: C\<^sub>0_def) lemma V\<^sub>0_power: "V\<^sub>0 v ^ n = monomial 1 (single_pp v n)" by (induction n) (auto simp: V\<^sub>0_def mult_single single_pp_plus) lemma single_MP_oalist [code]: "Poly_Mapping.single k v = distr\<^sub>0 nat_term_order_of_le [(k, v)]" unfolding distr\<^sub>0_def by (rule poly_mapping_eqI, simp add: lookup_single OAlist_lookup_ntm_single) lemma one_MP_oalist [code]: "1 = distr\<^sub>0 nat_term_order_of_le [(0, 1)]" by (metis single_MP_oalist single_one) lemma except_MP_oalist [code]: "except (MP_oalist xs) S = MP_oalist (OAlist_filter_ntm (\kv. fst kv \ S) xs)" by (rule poly_mapping_eqI, simp add: lookup_except oa_ntm.lookup_filter) subsubsection \Changing the Internal Order\ definition change_ord :: "'a::nat_term_compare nat_term_order \ ('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b)" where "change_ord to = (\x. x)" lemma change_ord_MP_oalist [code]: "change_ord to (MP_oalist xs) = MP_oalist (OAlist_reorder_ntm to xs)" by (rule poly_mapping_eqI, simp add: change_ord_def oa_ntm.lookup_reorder) subsubsection \Ordered Power-Products\ lemma foldl_assoc: assumes "\x y z. f (f x y) z = f x (f y z)" shows "foldl f (f a b) xs = f a (foldl f b xs)" proof (induct xs arbitrary: a b) fix a b show "foldl f (f a b) [] = f a (foldl f b [])" by simp next fix a b x xs assume "\a b. foldl f (f a b) xs = f a (foldl f b xs)" from assms[of a b x] this[of a "f b x"] show "foldl f (f a b) (x # xs) = f a (foldl f b (x # xs))" unfolding foldl_Cons by simp qed context gd_nat_term begin definition ord_pp :: "'a \ 'a \ bool" where "ord_pp s t = le_of_nat_term_order cmp_term (term_of_pair (s, the_min)) (term_of_pair (t, the_min))" definition ord_pp_strict :: "'a \ 'a \ bool" where "ord_pp_strict s t = lt_of_nat_term_order cmp_term (term_of_pair (s, the_min)) (term_of_pair (t, the_min))" lemma lt_MP_oalist [code]: "lt (MP_oalist xs) = (if is_zero (MP_oalist xs) then min_term else fst (OAlist_min_key_val_ntm cmp_term xs))" proof (split if_split, intro conjI impI) assume "is_zero (MP_oalist xs)" thus "lt (MP_oalist xs) = min_term" by (simp add: is_zero_def) next assume "\ is_zero (MP_oalist xs)" hence "fst (list_of_oalist_ntm xs) \ []" by (simp add: is_zero_MP_oalist List.null_def) show "lt (MP_oalist xs) = fst (OAlist_min_key_val_ntm cmp_term xs)" proof (rule lt_eqI_keys) show "fst (OAlist_min_key_val_ntm cmp_term xs) \ keys (MP_oalist xs)" by (simp add: keys_MP_oalist, rule imageI, rule oa_ntm.min_key_val_in, fact) next fix u assume "u \ keys (MP_oalist xs)" also have "... = fst ` set (fst (list_of_oalist_ntm xs))" by (simp add: keys_MP_oalist) finally obtain z where "z \ set (fst (list_of_oalist_ntm xs))" and "u = fst z" .. from this(1) have "ko.le (key_order_of_nat_term_order_inv cmp_term) (fst (OAlist_min_key_val_ntm cmp_term xs)) u" unfolding \u = fst z\ by (rule oa_ntm.min_key_val_minimal) thus "le_of_nat_term_order cmp_term u (fst (OAlist_min_key_val_ntm cmp_term xs))" by (simp add: le_of_nat_term_order_alt) qed qed lemma lc_MP_oalist [code]: "lc (MP_oalist xs) = (if is_zero (MP_oalist xs) then 0 else snd (OAlist_min_key_val_ntm cmp_term xs))" proof (split if_split, intro conjI impI) assume "is_zero (MP_oalist xs)" thus "lc (MP_oalist xs) = 0" by (simp add: is_zero_def) next assume "\ is_zero (MP_oalist xs)" moreover from this have "fst (list_of_oalist_ntm xs) \ []" by (simp add: is_zero_MP_oalist List.null_def) ultimately show "lc (MP_oalist xs) = snd (OAlist_min_key_val_ntm cmp_term xs)" by (simp add: lc_def lt_MP_oalist oa_ntm.snd_min_key_val) qed lemma tail_MP_oalist [code]: "tail (MP_oalist xs) = MP_oalist (OAlist_except_min_ntm cmp_term xs)" proof (cases "is_zero (MP_oalist xs)") case True hence "fst (list_of_oalist_ntm xs) = []" by (simp add: is_zero_MP_oalist List.null_def) hence "fst (list_of_oalist_ntm (OAlist_except_min_ntm cmp_term xs)) = []" by (rule oa_ntm.except_min_Nil) hence "is_zero (MP_oalist (OAlist_except_min_ntm cmp_term xs))" by (simp add: is_zero_MP_oalist List.null_def) with True show ?thesis by (simp add: is_zero_def) next case False show ?thesis by (rule poly_mapping_eqI, simp add: lookup_tail_2 oa_ntm.lookup_except_min' lt_MP_oalist False) qed definition comp_opt_p :: "('t \\<^sub>0 'c::zero, 't \\<^sub>0 'c) comp_opt" where "comp_opt_p p q = (if p = q then Some Eq else if ord_strict_p p q then Some Lt else if ord_strict_p q p then Some Gt else None)" lemma comp_opt_p_MP_oalist [code]: "comp_opt_p (MP_oalist xs) (MP_oalist ys) = OAlist_lex_ord_ntm cmp_term (\_ x y. if x = y then Some Eq else if x = 0 then Some Lt else if y = 0 then Some Gt else None) xs ys" proof - let ?f = "\_ x y. if x = y then Some Eq else if x = 0 then Some Lt else if y = 0 then Some Gt else None" show ?thesis proof (cases "comp_opt_p (MP_oalist xs) (MP_oalist ys) = Some Eq") case True hence "MP_oalist xs = MP_oalist ys" by (simp add: comp_opt_p_def split: if_splits) hence "lookup (MP_oalist xs) = lookup (MP_oalist ys)" by (rule arg_cong) hence eq: "OAlist_lookup_ntm xs = OAlist_lookup_ntm ys" by simp have "OAlist_lex_ord_ntm cmp_term ?f xs ys = Some Eq" by (rule oa_ntm.lex_ord_EqI, simp add: eq) with True show ?thesis by simp next case False hence neq: "MP_oalist xs \ MP_oalist ys" by (simp add: comp_opt_p_def split: if_splits) then obtain v where 1: "v \ keys (MP_oalist xs) \ keys (MP_oalist ys)" and 2: "lookup (MP_oalist xs) v \ lookup (MP_oalist ys) v" and 3: "\u. lt_of_nat_term_order cmp_term v u \ lookup (MP_oalist xs) u = lookup (MP_oalist ys) u" by (rule poly_mapping_neqE, blast) show ?thesis proof (rule HOL.sym, rule oa_ntm.lex_ord_valI) from 1 show "v \ fst ` set (fst (list_of_oalist_ntm xs)) \ fst ` set (fst (list_of_oalist_ntm ys))" by (simp add: keys_MP_oalist) next from 2 have 4: "OAlist_lookup_ntm xs v \ OAlist_lookup_ntm ys v" by simp show "comp_opt_p (MP_oalist xs) (MP_oalist ys) = (if OAlist_lookup_ntm xs v = OAlist_lookup_ntm ys v then Some Eq else if OAlist_lookup_ntm xs v = 0 then Some Lt else if OAlist_lookup_ntm ys v = 0 then Some Gt else None)" proof (simp add: 4, intro conjI impI) assume "OAlist_lookup_ntm ys v = 0" and "OAlist_lookup_ntm xs v = 0" with 4 show "comp_opt_p (MP_oalist xs) (MP_oalist ys) = Some Lt" by simp next assume "OAlist_lookup_ntm xs v \ 0" and "OAlist_lookup_ntm ys v = 0" hence "lookup (MP_oalist ys) v = 0" and "lookup (MP_oalist xs) v \ 0" by simp_all hence "ord_strict_p (MP_oalist ys) (MP_oalist xs)" using 3[symmetric] by (rule ord_strict_pI) with neq show "comp_opt_p (MP_oalist xs) (MP_oalist ys) = Some Gt" by (auto simp: comp_opt_p_def) next assume "OAlist_lookup_ntm ys v \ 0" and "OAlist_lookup_ntm xs v = 0" hence "lookup (MP_oalist xs) v = 0" and "lookup (MP_oalist ys) v \ 0" by simp_all hence "ord_strict_p (MP_oalist xs) (MP_oalist ys)" using 3 by (rule ord_strict_pI) with neq show "comp_opt_p (MP_oalist xs) (MP_oalist ys) = Some Lt" by (auto simp: comp_opt_p_def) next assume "OAlist_lookup_ntm xs v \ 0" hence "lookup (MP_oalist xs) v \ 0" by simp with 2 have a: "\ ord_strict_p (MP_oalist xs) (MP_oalist ys)" using 3 by (rule not_ord_strict_pI) assume "OAlist_lookup_ntm ys v \ 0" hence "lookup (MP_oalist ys) v \ 0" by simp with 2[symmetric] have "\ ord_strict_p (MP_oalist ys) (MP_oalist xs)" using 3[symmetric] by (rule not_ord_strict_pI) with neq a show "comp_opt_p (MP_oalist xs) (MP_oalist ys) = None" by (auto simp: comp_opt_p_def) qed next fix u assume "ko.lt (key_order_of_nat_term_order_inv cmp_term) u v" hence "lt_of_nat_term_order cmp_term v u" by (simp only: lt_of_nat_term_order_alt) hence "lookup (MP_oalist xs) u = lookup (MP_oalist ys) u" by (rule 3) thus "(if OAlist_lookup_ntm xs u = OAlist_lookup_ntm ys u then Some Eq else if OAlist_lookup_ntm xs u = 0 then Some Lt else if OAlist_lookup_ntm ys u = 0 then Some Gt else None) = Some Eq" by simp qed fact qed qed lemma compute_ord_p [code]: "ord_p p q = (let aux = comp_opt_p p q in aux = Some Lt \ aux = Some Eq)" by (auto simp: ord_p_def comp_opt_p_def) lemma compute_ord_p_strict [code]: "ord_strict_p p q = (comp_opt_p p q = Some Lt)" by (auto simp: comp_opt_p_def) lemma keys_to_list_MP_oalist [code]: "keys_to_list (MP_oalist xs) = OAlist_sorted_domain_ntm cmp_term xs" proof - have eq: "ko.lt (key_order_of_nat_term_order_inv cmp_term) = ord_term_strict_conv" by (intro ext, simp add: lt_of_nat_term_order_alt) have 1: "irreflp ord_term_strict_conv" by (rule irreflpI, simp) have 2: "transp ord_term_strict_conv" by (rule transpI, simp) have "antisymp ord_term_strict_conv" by (rule antisympI, simp) moreover have 3: "sorted_wrt ord_term_strict_conv (keys_to_list (MP_oalist xs))" unfolding keys_to_list_def by (fact pps_to_list_sorted_wrt) moreover note _ moreover have 4: "sorted_wrt ord_term_strict_conv (OAlist_sorted_domain_ntm cmp_term xs)" unfolding eq[symmetric] by (fact oa_ntm.sorted_sorted_domain) ultimately show ?thesis proof (rule sorted_wrt_distinct_set_unique) from 1 2 3 show "distinct (keys_to_list (MP_oalist xs))" by (rule distinct_sorted_wrt_irrefl) next from 1 2 4 show "distinct (OAlist_sorted_domain_ntm cmp_term xs)" by (rule distinct_sorted_wrt_irrefl) next show "set (keys_to_list (MP_oalist xs)) = set (OAlist_sorted_domain_ntm cmp_term xs)" by (simp add: set_keys_to_list keys_MP_oalist oa_ntm.set_sorted_domain) qed qed end (* ordered_nat_term *) lifting_update poly_mapping.lifting lifting_forget poly_mapping.lifting subsection \Interpretations\ lemma term_powerprod_gd_term: fixes pair_of_term :: "'t::nat_term \ ('a::{graded_dickson_powerprod,nat_pp_compare} \ 'k::{the_min,wellorder})" assumes "term_powerprod pair_of_term term_of_pair" and "\v. fst (rep_nat_term v) = rep_nat_pp (fst (pair_of_term v))" and "\t. snd (rep_nat_term (term_of_pair (t, the_min))) = 0" and "\v w. snd (pair_of_term v) \ snd (pair_of_term w) \ snd (rep_nat_term v) \ snd (rep_nat_term w)" and "\s t k. term_of_pair (s + t, k) = splus (term_of_pair (s, k)) (term_of_pair (t, k))" and "\t v. term_powerprod.splus pair_of_term term_of_pair t v = splus (term_of_pair (t, the_min)) v" shows "gd_term pair_of_term term_of_pair (\s t. le_of_nat_term_order cmp_term (term_of_pair (s, the_min)) (term_of_pair (t, the_min))) (\s t. lt_of_nat_term_order cmp_term (term_of_pair (s, the_min)) (term_of_pair (t, the_min))) (le_of_nat_term_order cmp_term) (lt_of_nat_term_order cmp_term)" proof - from assms(1) interpret tp: term_powerprod pair_of_term term_of_pair . let ?f = "\x. term_of_pair (x, the_min)" show ?thesis proof (intro gd_term.intro ordered_term.intro) from assms(1) show "term_powerprod pair_of_term term_of_pair" . next show "ordered_powerprod (\s t. le_of_nat_term_order cmp_term (?f s) (?f t)) (\s t. lt_of_nat_term_order cmp_term (?f s) (?f t))" proof (intro ordered_powerprod.intro ordered_powerprod_axioms.intro) show "class.linorder (\s t. le_of_nat_term_order cmp_term (?f s) (?f t)) (\s t. lt_of_nat_term_order cmp_term (?f s) (?f t))" proof (unfold_locales, simp_all add: lt_of_nat_term_order_alt le_of_nat_term_order_alt ko.linear ko.less_le_not_le) fix x y assume "ko.le (key_order_of_nat_term_order_inv cmp_term) (term_of_pair (x, the_min)) (term_of_pair (y, the_min))" and "ko.le (key_order_of_nat_term_order_inv cmp_term) (term_of_pair (y, the_min)) (term_of_pair (x, the_min))" hence "term_of_pair (x, the_min) = term_of_pair (y, the_min)" by (rule ko.antisym) hence "(x, the_min) = (y, the_min::'k)" by (rule tp.term_of_pair_injective) thus "x = y" by simp qed next fix t show "le_of_nat_term_order cmp_term (?f 0) (?f t)" unfolding le_of_nat_term_order by (rule nat_term_compD1', fact comparator_nat_term_compare, fact nat_term_comp_nat_term_compare, simp add: assms(3), simp add: assms(2) zero_pp tp.pair_term) next fix s t u assume "le_of_nat_term_order cmp_term (?f s) (?f t)" hence "le_of_nat_term_order cmp_term (?f (u + s)) (?f (u + t))" by (simp add: le_of_nat_term_order assms(5) nat_term_compare_splus) thus "le_of_nat_term_order cmp_term (?f (s + u)) (?f (t + u))" by (simp only: ac_simps) qed next show "class.linorder (le_of_nat_term_order cmp_term) (lt_of_nat_term_order cmp_term)" by (fact linorder_le_of_nat_term_order) next show "ordered_term_axioms pair_of_term term_of_pair (\s t. le_of_nat_term_order cmp_term (?f s) (?f t)) (le_of_nat_term_order cmp_term)" proof fix v w t assume "le_of_nat_term_order cmp_term v w" thus "le_of_nat_term_order cmp_term (t \ v) (t \ w)" by (simp add: le_of_nat_term_order assms(6) nat_term_compare_splus) next fix v w assume "le_of_nat_term_order cmp_term (?f (tp.pp_of_term v)) (?f (tp.pp_of_term w))" hence 3: "nat_term_compare cmp_term (?f (tp.pp_of_term v)) (?f (tp.pp_of_term w)) \ Gt" by (simp add: le_of_nat_term_order) assume "tp.component_of_term v \ tp.component_of_term w" hence 4: "snd (rep_nat_term v) \ snd (rep_nat_term w)" by (simp add: tp.component_of_term_def assms(4)) note comparator_nat_term_compare nat_term_comp_nat_term_compare moreover have "fst (rep_nat_term v) = fst (rep_nat_term (?f (tp.pp_of_term v)))" by (simp add: assms(2) tp.pp_of_term_def tp.pair_term) moreover have "fst (rep_nat_term w) = fst (rep_nat_term (?f (tp.pp_of_term w)))" by (simp add: assms(2) tp.pp_of_term_def tp.pair_term) moreover note 4 moreover have "snd (rep_nat_term (?f (tp.pp_of_term v))) = snd (rep_nat_term (?f (tp.pp_of_term w)))" by (simp add: assms(3)) ultimately show "le_of_nat_term_order cmp_term v w" unfolding le_of_nat_term_order using 3 by (rule nat_term_compD4'') qed qed qed lemma gd_term_to_pair_unit: "gd_term (to_pair_unit::'a::{nat_term_compare,nat_pp_term,graded_dickson_powerprod} \ _) fst (\s t. le_of_nat_term_order cmp_term (fst (s, the_min)) (fst (t, the_min))) (\s t. lt_of_nat_term_order cmp_term (fst (s, the_min)) (fst (t, the_min))) (le_of_nat_term_order cmp_term) (lt_of_nat_term_order cmp_term)" proof (intro gd_term.intro ordered_term.intro) show "term_powerprod to_pair_unit fst" by unfold_locales next show "ordered_powerprod (\s t. le_of_nat_term_order cmp_term (fst (s, the_min)) (fst (t, the_min))) (\s t. lt_of_nat_term_order cmp_term (fst (s, the_min)) (fst (t, the_min)))" unfolding fst_conv using linorder_le_of_nat_term_order proof (intro ordered_powerprod.intro) from le_of_nat_term_order_zero_min show "ordered_powerprod_axioms (le_of_nat_term_order cmp_term)" proof (unfold_locales) fix s t u assume "le_of_nat_term_order cmp_term s t" hence "le_of_nat_term_order cmp_term (u + s) (u + t)" by (rule le_of_nat_term_order_plus_monotone) thus "le_of_nat_term_order cmp_term (s + u) (t + u)" by (simp only: ac_simps) qed qed next show "class.linorder (le_of_nat_term_order cmp_term) (lt_of_nat_term_order cmp_term)" by (fact linorder_le_of_nat_term_order) next show "ordered_term_axioms to_pair_unit fst (\s t. le_of_nat_term_order cmp_term (fst (s, the_min)) (fst (t, the_min))) (le_of_nat_term_order cmp_term)" by (unfold_locales, auto intro: le_of_nat_term_order_plus_monotone) qed corollary gd_nat_term_to_pair_unit: "gd_nat_term (to_pair_unit::'a::{nat_term_compare,nat_pp_term,graded_dickson_powerprod} \ _) fst cmp_term" by (rule gd_nat_term.intro, fact gd_term_to_pair_unit, rule gd_nat_term_axioms.intro, simp add: splus_pp_term) lemma gd_term_id: "gd_term (\x::('a::{nat_term_compare,nat_pp_compare,nat_pp_term,graded_dickson_powerprod} \ 'b::{nat,the_min}). x) (\x. x) (\s t. le_of_nat_term_order cmp_term (s, the_min) (t, the_min)) (\s t. lt_of_nat_term_order cmp_term (s, the_min) (t, the_min)) (le_of_nat_term_order cmp_term) (lt_of_nat_term_order cmp_term)" apply (rule term_powerprod_gd_term) subgoal by unfold_locales subgoal by (simp add: rep_nat_term_prod_def) subgoal by (simp add: rep_nat_term_prod_def the_min_eq_zero) subgoal by (simp add: rep_nat_term_prod_def ord_iff[symmetric]) subgoal by (simp add: splus_prod_def pprod.splus_def) subgoal by (simp add: splus_prod_def) done corollary gd_nat_term_id: "gd_nat_term (\x. x) (\x. x) cmp_term" for cmp_term :: "('a::{nat_term_compare,nat_pp_compare,nat_pp_term,graded_dickson_powerprod} \ 'c::{nat,the_min}) nat_term_order" by (rule gd_nat_term.intro, fact gd_term_id, rule gd_nat_term_axioms.intro, simp add: splus_prod_def) subsection \Computations\ type_synonym 'a mpoly_tc = "(nat, nat) pp \\<^sub>0 'a" global_interpretation punit0: gd_nat_term "to_pair_unit::'a::{nat_term_compare,nat_pp_term,graded_dickson_powerprod} \ _" fst cmp_term rewrites "punit.adds_term = (adds)" and "punit.pp_of_term = (\x. x)" and "punit.component_of_term = (\_. ())" for cmp_term defines monom_mult_punit = punit.monom_mult and mult_scalar_punit = punit.mult_scalar and shift_map_keys_punit = punit0.shift_map_keys and ord_pp_punit = punit0.ord_pp and ord_pp_strict_punit = punit0.ord_pp_strict and min_term_punit = punit0.min_term and lt_punit = punit0.lt and lc_punit = punit0.lc and tail_punit = punit0.tail and comp_opt_p_punit = punit0.comp_opt_p and ord_p_punit = punit0.ord_p and ord_strict_p_punit = punit0.ord_strict_p and keys_to_list_punit = punit0.keys_to_list subgoal by (fact gd_nat_term_to_pair_unit) subgoal by (fact punit_adds_term) subgoal by (fact punit_pp_of_term) subgoal by (fact punit_component_of_term) done lemma shift_map_keys_punit_MP_oalist [code abstract]: "list_of_oalist_ntm (shift_map_keys_punit t f xs) = map_raw (\(k, v). (t + k, f v)) (list_of_oalist_ntm xs)" by (simp add: punit0.list_of_oalist_shift_keys case_prod_beta') lemmas [code] = punit0.mult_scalar_MP_oalist[unfolded mult_scalar_punit_def punit_mult_scalar] punit0.punit_min_term lemma ord_pp_punit_alt [code_unfold]: "ord_pp_punit = le_of_nat_term_order" by (intro ext, simp add: punit0.ord_pp_def) lemma ord_pp_strict_punit_alt [code_unfold]: "ord_pp_strict_punit = lt_of_nat_term_order" by (intro ext, simp add: punit0.ord_pp_strict_def) lemma gd_powerprod_ord_pp_punit: "gd_powerprod (ord_pp_punit cmp_term) (ord_pp_strict_punit cmp_term)" unfolding punit0.ord_pp_def punit0.ord_pp_strict_def .. locale trivariate\<^sub>0_rat begin abbreviation X::"rat mpoly_tc" where "X \ V\<^sub>0 (0::nat)" abbreviation Y::"rat mpoly_tc" where "Y \ V\<^sub>0 (1::nat)" abbreviation Z::"rat mpoly_tc" where "Z \ V\<^sub>0 (2::nat)" end experiment begin interpretation trivariate\<^sub>0_rat . value [code] "X ^ 2" value [code] "X\<^sup>2 * Z + 2 * Y ^ 3 * Z\<^sup>2" value [code] "distr\<^sub>0 DRLEX [(sparse\<^sub>0 [(0::nat, 3::nat)], 1::rat)] = distr\<^sub>0 DRLEX [(sparse\<^sub>0 [(0, 3)], 1)]" lemma "ord_strict_p_punit DRLEX (X\<^sup>2 * Z + 2 * Y ^ 3 * Z\<^sup>2) (X\<^sup>2 * Z\<^sup>2 + 2 * Y ^ 3 * Z\<^sup>2)" by eval lemma "tail_punit DLEX (X\<^sup>2 * Z + 2 * Y ^ 3 * Z\<^sup>2) = X\<^sup>2 * Z" by eval value [code] "min_term_punit::(nat, nat) pp" value [code] "is_zero (distr\<^sub>0 DRLEX [(sparse\<^sub>0 [(0::nat, 3::nat)], 1::rat)])" value [code] "lt_punit DRLEX (distr\<^sub>0 DRLEX [(sparse\<^sub>0 [(0::nat, 3::nat)], 1::rat)])" lemma "lt_punit DRLEX (X\<^sup>2 * Z + 2 * Y ^ 3 * Z\<^sup>2) = sparse\<^sub>0 [(1, 3), (2, 2)]" by eval lemma "lt_punit DRLEX (X + Y + Z) = sparse\<^sub>0 [(2, 1)]" by eval lemma "keys (X\<^sup>2 * Z ^ 3 + 2 * Y ^ 3 * Z\<^sup>2) = {sparse\<^sub>0 [(0, 2), (2, 3)], sparse\<^sub>0 [(1, 3), (2, 2)]}" by eval lemma "- 1 * X\<^sup>2 * Z ^ 7 + - 2 * Y ^ 3 * Z\<^sup>2 = - X\<^sup>2 * Z ^ 7 + - 2 * Y ^ 3 * Z\<^sup>2" by eval lemma "X\<^sup>2 * Z ^ 7 + 2 * Y ^ 3 * Z\<^sup>2 + X\<^sup>2 * Z ^ 4 + - 2 * Y ^ 3 * Z\<^sup>2 = X\<^sup>2 * Z ^ 7 + X\<^sup>2 * Z ^ 4" by eval lemma "X\<^sup>2 * Z ^ 7 + 2 * Y ^ 3 * Z\<^sup>2 - X\<^sup>2 * Z ^ 4 + - 2 * Y ^ 3 * Z\<^sup>2 = X\<^sup>2 * Z ^ 7 - X\<^sup>2 * Z ^ 4" by eval lemma "lookup (X\<^sup>2 * Z ^ 7 + 2 * Y ^ 3 * Z\<^sup>2 + 2) (sparse\<^sub>0 [(0, 2), (2, 7)]) = 1" by eval lemma "X\<^sup>2 * Z ^ 7 + 2 * Y ^ 3 * Z\<^sup>2 \ X\<^sup>2 * Z ^ 4 + - 2 * Y ^ 3 * Z\<^sup>2" by eval lemma "0 * X^2 * Z^7 + 0 * Y^3*Z\<^sup>2 = 0" by eval lemma "monom_mult_punit 3 (sparse\<^sub>0 [(1, 2::nat)]) (X\<^sup>2 * Z + 2 * Y ^ 3 * Z\<^sup>2) = 3 * Y\<^sup>2 * Z * X\<^sup>2 + 6 * Y ^ 5 * Z\<^sup>2" by eval lemma "monomial (-4) (sparse\<^sub>0 [(0, 2::nat)]) = - 4 * X\<^sup>2" by eval lemma "monomial (0::rat) (sparse\<^sub>0 [(0::nat, 2::nat)]) = 0" by eval lemma "(X\<^sup>2 * Z + 2 * Y ^ 3 * Z\<^sup>2) * (X\<^sup>2 * Z ^ 3 + - 2 * Y ^ 3 * Z\<^sup>2) = X ^ 4 * Z ^ 4 + - 2 * X\<^sup>2 * Z ^ 3 * Y ^ 3 + - 4 * Y ^ 6 * Z ^ 4 + 2 * Y ^ 3 * Z ^ 5 * X\<^sup>2" by eval end subsection \Code setup for type MPoly\ text \postprocessing from \Var\<^sub>0, Const\<^sub>0\ to \Var, Const\.\ lemmas [code_post] = plus_mpoly.abs_eq[symmetric] times_mpoly.abs_eq[symmetric] one_mpoly_def[symmetric] Var.abs_eq[symmetric] Const.abs_eq[symmetric] instantiation mpoly::("{equal, zero}")equal begin lift_definition equal_mpoly:: "'a mpoly \ 'a mpoly \ bool" is HOL.equal . instance proof standard qed (transfer, rule equal_eq) end end (* theory *) diff --git a/thys/Polynomials/OAlist.thy b/thys/Polynomials/OAlist.thy --- a/thys/Polynomials/OAlist.thy +++ b/thys/Polynomials/OAlist.thy @@ -1,4282 +1,4282 @@ (* Author: Florian Haftmann, TU Muenchen *) (* Author: Andreas Lochbihler, ETH Zurich *) (* Author: Alexander Maletzky, RISC Linz *) section \Associative Lists with Sorted Keys\ theory OAlist imports Deriving.Comparator begin text \We define the type of @{emph \ordered associative lists\} (oalist). An oalist is an associative list (i.\,e. a list of pairs) such that the keys are distinct and sorted wrt. some linear order relation, and no key is mapped to @{term 0}. The latter invariant allows to implement various functions operating on oalists more efficiently. The ordering of the keys in an oalist \xs\ is encoded as an additional parameter of \xs\. This means that oalists may be ordered wrt. different orderings, even if they are of the same type. Operations operating on more than one oalists, like \map2_val\, typically ensure that the orderings of their arguments are identical by re-ordering one argument wrt. the order relation of the other. This, however, implies that equality of order relations must be effectively decidable if executable code is to be generated.\ subsection \Preliminaries\ fun min_list_param :: "('a \ 'a \ bool) \ 'a list \ 'a" where "min_list_param rel (x # xs) = (case xs of [] \ x | _ \ (let m = min_list_param rel xs in if rel x m then x else m))" lemma min_list_param_in: assumes "xs \ []" shows "min_list_param rel xs \ set xs" using assms proof (induct xs) case Nil thus ?case by simp next case (Cons x xs) show ?case proof (simp add: min_list_param.simps[of rel x xs] Let_def del: min_list_param.simps set_simps(2) split: list.split, intro conjI impI allI, simp, simp) fix y ys assume xs: "xs = y # ys" have "min_list_param rel (y # ys) = min_list_param rel xs" by (simp only: xs) also have "... \ set xs" by (rule Cons(1), simp add: xs) also have "... \ set (x # y # ys)" by (auto simp: xs) finally show "min_list_param rel (y # ys) \ set (x # y # ys)" . qed qed lemma min_list_param_minimal: assumes "transp rel" and "\x y. x \ set xs \ y \ set xs \ rel x y \ rel y x" and "z \ set xs" shows "rel (min_list_param rel xs) z" using assms(2, 3) proof (induct xs) case Nil from Nil(2) show ?case by simp next case (Cons x xs) from Cons(3) have disj1: "z = x \ z \ set xs" by simp have "x \ set (x # xs)" by simp hence disj2: "rel x z \ rel z x" using Cons(3) by (rule Cons(2)) have *: "rel (min_list_param rel xs) z" if "z \ set xs" using _ that proof (rule Cons(1)) fix a b assume "a \ set xs" and "b \ set xs" hence "a \ set (x # xs)" and "b \ set (x # xs)" by simp_all thus "rel a b \ rel b a" by (rule Cons(2)) qed show ?case proof (simp add: min_list_param.simps[of rel x xs] Let_def del: min_list_param.simps set_simps(2) split: list.split, intro conjI impI allI) assume "xs = []" with disj1 disj2 show "rel x z" by simp next fix y ys assume "xs = y # ys" and "rel x (min_list_param rel (y # ys))" hence "rel x (min_list_param rel xs)" by simp from disj1 show "rel x z" proof assume "z = x" thus ?thesis using disj2 by simp next assume "z \ set xs" hence "rel (min_list_param rel xs) z" by (rule *) with assms(1) \rel x (min_list_param rel xs)\ show ?thesis by (rule transpD) qed next fix y ys assume xs: "xs = y # ys" and "\ rel x (min_list_param rel (y # ys))" from disj1 show "rel (min_list_param rel (y # ys)) z" proof assume "z = x" have "min_list_param rel (y # ys) \ set (y # ys)" by (rule min_list_param_in, simp) hence "min_list_param rel (y # ys) \ set (x # xs)" by (simp add: xs) with \x \ set (x # xs)\ have "rel x (min_list_param rel (y # ys)) \ rel (min_list_param rel (y # ys)) x" by (rule Cons(2)) with \\ rel x (min_list_param rel (y # ys))\ have "rel (min_list_param rel (y # ys)) x" by simp thus ?thesis by (simp only: \z = x\) next assume "z \ set xs" hence "rel (min_list_param rel xs) z" by (rule *) thus ?thesis by (simp only: xs) qed qed qed definition comp_of_ord :: "('a \ 'a \ bool) \ 'a comparator" where "comp_of_ord le x y = (if le x y then if x = y then Eq else Lt else Gt)" lemma comp_of_ord_eq_comp_of_ords: assumes "antisymp le" shows "comp_of_ord le = comp_of_ords le (\x y. le x y \ \ le y x)" by (intro ext, auto simp: comp_of_ord_def comp_of_ords_def intro: assms antisympD) lemma comparator_converse: assumes "comparator cmp" shows "comparator (\x y. cmp y x)" proof - from assms interpret comp?: comparator cmp . show ?thesis by (unfold_locales, auto simp: comp.eq comp.sym intro: comp_trans) qed lemma comparator_composition: assumes "comparator cmp" and "inj f" shows "comparator (\x y. cmp (f x) (f y))" proof - from assms(1) interpret comp?: comparator cmp . from assms(2) have *: "x = y" if "f x = f y" for x y using that by (rule injD) show ?thesis by (unfold_locales, auto simp: comp.sym comp.eq * intro: comp_trans) qed (* subsection \Syntactic Type Class for Default Elements\ text \We do not want to use the existing type-class @{class default}, but instead introduce a new one:\ class oalist_dflt = fixes dflt::'a simproc_setup reorient_dflt ("dflt = x") = Reorient_Proc.proc *) subsection \Type \key_order\\ typedef 'a key_order = "{compare :: 'a comparator. comparator compare}" morphisms key_compare Abs_key_order proof - from well_order_on obtain r where "well_order_on (UNIV::'a set) r" .. hence "linear_order r" by (simp only: well_order_on_def) hence lin: "(x, y) \ r \ (y, x) \ r" for x y by (metis Diff_iff Linear_order_in_diff_Id UNIV_I \well_order r\ well_order_on_Field) have antisym: "(x, y) \ r \ (y, x) \ r \ x = y" for x y by (meson \linear_order r\ antisymD linear_order_on_def partial_order_on_def) have trans: "(x, y) \ r \ (y, z) \ r \ (x, z) \ r" for x y z by (meson \linear_order r\ linear_order_on_def order_on_defs(1) partial_order_on_def trans_def) define comp where "comp = (\x y. if (x, y) \ r then if (y, x) \ r then Eq else Lt else Gt)" show ?thesis proof (rule, simp) show "comparator comp" proof (standard, simp_all add: comp_def split: if_splits, intro impI) fix x y assume "(x, y) \ r" and "(y, x) \ r" thus "x = y" by (rule antisym) next fix x y assume "(x, y) \ r" with lin show "(y, x) \ r" by blast next fix x y z assume "(y, x) \ r" and "(z, y) \ r" assume "(x, y) \ r" and "(y, z) \ r" hence "(x, z) \ r" by (rule trans) moreover have "(z, x) \ r" proof assume "(z, x) \ r" with \(x, z) \ r\ have "x = z" by (rule antisym) from \(z, y) \ r\ \(x, y) \ r\ show False unfolding \x = z\ .. qed ultimately show "(z, x) \ r \ ((z, x) \ r \ (x, z) \ r)" by simp qed qed qed lemma comparator_key_compare [simp, intro!]: "comparator (key_compare ko)" using key_compare[of ko] by simp instantiation key_order :: (type) equal begin definition equal_key_order :: "'a key_order \ 'a key_order \ bool" where "equal_key_order = (=)" instance by (standard, simp add: equal_key_order_def) end setup_lifting type_definition_key_order instantiation key_order :: (type) uminus begin lift_definition uminus_key_order :: "'a key_order \ 'a key_order" is "\c x y. c y x" by (fact comparator_converse) instance .. end lift_definition le_of_key_order :: "'a key_order \ 'a \ 'a \ bool" is "\cmp. le_of_comp cmp" . lift_definition lt_of_key_order :: "'a key_order \ 'a \ 'a \ bool" is "\cmp. lt_of_comp cmp" . definition key_order_of_ord :: "('a \ 'a \ bool) \ 'a key_order" where "key_order_of_ord ord = Abs_key_order (comp_of_ord ord)" lift_definition key_order_of_le :: "'a::linorder key_order" is comparator_of by (fact comparator_of) interpretation key_order_lin: linorder "le_of_key_order ko" "lt_of_key_order ko" proof transfer fix comp::"'a comparator" assume "comparator comp" then interpret comp: comparator comp . show "class.linorder comp.le comp.lt" by (fact comp.linorder) qed lemma le_of_key_order_alt: "le_of_key_order ko x y = (key_compare ko x y \ Gt)" by (transfer, simp add: comparator.nGt_le_conv) lemma lt_of_key_order_alt: "lt_of_key_order ko x y = (key_compare ko x y = Lt)" by (transfer, meson comparator.Lt_lt_conv) lemma key_compare_Gt: "key_compare ko x y = Gt \ key_compare ko y x = Lt" by (transfer, meson comparator.nGt_le_conv comparator.nLt_le_conv) lemma key_compare_Eq: "key_compare ko x y = Eq \ x = y" by (transfer, simp add: comparator.eq) lemma key_compare_same [simp]: "key_compare ko x x = Eq" by (simp add: key_compare_Eq) lemma uminus_key_compare [simp]: "invert_order (key_compare ko x y) = key_compare ko y x" by (transfer, simp add: comparator.sym) lemma key_compare_uminus [simp]: "key_compare (- ko) x y = key_compare ko y x" by (transfer, rule refl) lemma uminus_key_order_sameD: assumes "- ko = (ko::'a key_order)" shows "x = (y::'a)" proof (rule ccontr) assume "x \ y" hence "key_compare ko x y \ Eq" by (simp add: key_compare_Eq) hence "key_compare ko x y \ invert_order (key_compare ko x y)" by (metis invert_order.elims order.distinct(5)) also have "invert_order (key_compare ko x y) = key_compare (- ko) x y" by simp finally have "- ko \ ko" by (auto simp del: key_compare_uminus) thus False using assms .. qed lemma key_compare_key_order_of_ord: assumes "antisymp ord" and "transp ord" and "\x y. ord x y \ ord y x" shows "key_compare (key_order_of_ord ord) = (\x y. if ord x y then if x = y then Eq else Lt else Gt)" proof - have eq: "key_compare (key_order_of_ord ord) = comp_of_ord ord" unfolding key_order_of_ord_def comp_of_ord_eq_comp_of_ords[OF assms(1)] proof (rule Abs_key_order_inverse, simp, rule comp_of_ords, unfold_locales) fix x from assms(3) show "ord x x" by blast next fix x y z assume "ord x y" and "ord y z" with assms(2) show "ord x z" by (rule transpD) next fix x y assume "ord x y" and "ord y x" with assms(1) show "x = y" by (rule antisympD) qed (rule refl, rule assms(3)) have *: "x = y" if "ord x y" and "ord y x" for x y using assms(1) that by (rule antisympD) show ?thesis by (rule, rule, auto simp: eq comp_of_ord_def intro: *) qed lemma key_compare_key_order_of_le: "key_compare key_order_of_le = (\x y. if x < y then Lt else if x = y then Eq else Gt)" by (transfer, intro ext, fact comparator_of_def) subsection \Invariant in Context @{locale comparator}\ context comparator begin definition oalist_inv_raw :: "('a \ 'b::zero) list \ bool" where "oalist_inv_raw xs \ (0 \ snd ` set xs \ sorted_wrt lt (map fst xs))" lemma oalist_inv_rawI: assumes "0 \ snd ` set xs" and "sorted_wrt lt (map fst xs)" shows "oalist_inv_raw xs" unfolding oalist_inv_raw_def using assms unfolding fst_conv snd_conv by blast lemma oalist_inv_rawD1: assumes "oalist_inv_raw xs" shows "0 \ snd ` set xs" using assms unfolding oalist_inv_raw_def fst_conv by blast lemma oalist_inv_rawD2: assumes "oalist_inv_raw xs" shows "sorted_wrt lt (map fst xs)" using assms unfolding oalist_inv_raw_def fst_conv snd_conv by blast lemma oalist_inv_raw_Nil: "oalist_inv_raw []" by (simp add: oalist_inv_raw_def) lemma oalist_inv_raw_singleton: "oalist_inv_raw [(k, v)] \ (v \ 0)" by (auto simp: oalist_inv_raw_def) lemma oalist_inv_raw_ConsI: assumes "oalist_inv_raw xs" and "v \ 0" and "xs \ [] \ lt k (fst (hd xs))" shows "oalist_inv_raw ((k, v) # xs)" proof (rule oalist_inv_rawI) from assms(1) have "0 \ snd ` set xs" by (rule oalist_inv_rawD1) with assms(2) show "0 \ snd ` set ((k, v) # xs)" by simp next show "sorted_wrt lt (map fst ((k, v) # xs))" proof (cases "xs = []") case True thus ?thesis by simp next case False then obtain k' v' xs' where xs: "xs = (k', v') # xs'" by (metis list.exhaust prod.exhaust) from assms(3)[OF False] have "lt k k'" by (simp add: xs) moreover from assms(1) have "sorted_wrt lt (map fst xs)" by (rule oalist_inv_rawD2) ultimately show "sorted_wrt lt (map fst ((k, v) # xs))" - by (simp add: xs sorted_wrt2[OF transp_less] del: sorted_wrt.simps) + by (simp add: xs sorted_wrt2[OF transp_on_less] del: sorted_wrt.simps) qed qed lemma oalist_inv_raw_ConsD1: assumes "oalist_inv_raw (x # xs)" shows "oalist_inv_raw xs" proof (rule oalist_inv_rawI) from assms have "0 \ snd ` set (x # xs)" by (rule oalist_inv_rawD1) thus "0 \ snd ` set xs" by simp next from assms have "sorted_wrt lt (map fst (x # xs))" by (rule oalist_inv_rawD2) thus "sorted_wrt lt (map fst xs)" by simp qed lemma oalist_inv_raw_ConsD2: assumes "oalist_inv_raw ((k, v) # xs)" shows "v \ 0" proof - from assms have "0 \ snd ` set ((k, v) # xs)" by (rule oalist_inv_rawD1) thus ?thesis by auto qed lemma oalist_inv_raw_ConsD3: assumes "oalist_inv_raw ((k, v) # xs)" and "k' \ fst ` set xs" shows "lt k k'" proof - from assms(2) obtain x where "x \ set xs" and "k' = fst x" by fastforce from assms(1) have "sorted_wrt lt (map fst ((k, v) # xs))" by (rule oalist_inv_rawD2) hence "\x\set xs. lt k (fst x)" by simp hence "lt k (fst x)" using \x \ set xs\ .. thus ?thesis by (simp only: \k' = fst x\) qed lemma oalist_inv_raw_tl: assumes "oalist_inv_raw xs" shows "oalist_inv_raw (tl xs)" proof (rule oalist_inv_rawI) from assms have "0 \ snd ` set xs" by (rule oalist_inv_rawD1) thus "0 \ snd ` set (tl xs)" by (metis (no_types, lifting) image_iff list.set_sel(2) tl_Nil) next show "sorted_wrt lt (map fst (tl xs))" by (metis hd_Cons_tl oalist_inv_rawD2 oalist_inv_raw_ConsD1 assms tl_Nil) qed lemma oalist_inv_raw_filter: assumes "oalist_inv_raw xs" shows "oalist_inv_raw (filter P xs)" proof (rule oalist_inv_rawI) from assms have "0 \ snd ` set xs" by (rule oalist_inv_rawD1) thus "0 \ snd ` set (filter P xs)" by auto next from assms have "sorted_wrt lt (map fst xs)" by (rule oalist_inv_rawD2) thus "sorted_wrt lt (map fst (filter P xs))" by (induct xs, simp, simp) qed lemma oalist_inv_raw_map: assumes "oalist_inv_raw xs" and "\a. snd (f a) = 0 \ snd a = 0" and "\a b. comp (fst (f a)) (fst (f b)) = comp (fst a) (fst b)" shows "oalist_inv_raw (map f xs)" proof (rule oalist_inv_rawI) show "0 \ snd ` set (map f xs)" proof (simp, rule) assume "0 \ snd ` f ` set xs" then obtain a where "a \ set xs" and "snd (f a) = 0" by fastforce from this(2) have "snd a = 0" by (rule assms(2)) from assms(1) have "0 \ snd ` set xs" by (rule oalist_inv_rawD1) moreover from \a \ set xs\ have "0 \ snd ` set xs" by (simp add: \snd a = 0\[symmetric]) ultimately show False .. qed next from assms(1) have "sorted_wrt lt (map fst xs)" by (rule oalist_inv_rawD2) hence "sorted_wrt (\x y. comp (fst x) (fst y) = Lt) xs" by (simp only: sorted_wrt_map Lt_lt_conv) thus "sorted_wrt lt (map fst (map f xs))" by (simp add: sorted_wrt_map Lt_lt_conv[symmetric] assms(3)) qed lemma oalist_inv_raw_induct [consumes 1, case_names Nil Cons]: assumes "oalist_inv_raw xs" assumes "P []" assumes "\k v xs. oalist_inv_raw ((k, v) # xs) \ oalist_inv_raw xs \ v \ 0 \ (\k'. k' \ fst ` set xs \ lt k k') \ P xs \ P ((k, v) # xs)" shows "P xs" using assms(1) proof (induct xs) case Nil from assms(2) show ?case . next case (Cons x xs) obtain k v where x: "x = (k, v)" by fastforce from Cons(2) have "oalist_inv_raw ((k, v) # xs)" and "oalist_inv_raw xs" and "v \ 0" unfolding x by (this, rule oalist_inv_raw_ConsD1, rule oalist_inv_raw_ConsD2) moreover from Cons(2) have "lt k k'" if "k' \ fst ` set xs" for k' using that unfolding x by (rule oalist_inv_raw_ConsD3) moreover from \oalist_inv_raw xs\ have "P xs" by (rule Cons(1)) ultimately show ?case unfolding x by (rule assms(3)) qed subsection \Operations on Lists of Pairs in Context @{locale comparator}\ type_synonym (in -) ('a, 'b) comp_opt = "'a \ 'b \ (order option)" definition (in -) lookup_dflt :: "('a \ 'b) list \ 'a \ 'b::zero" where "lookup_dflt xs k = (case map_of xs k of Some v \ v | None \ 0)" text \@{const lookup_dflt} is only an auxiliary function needed for proving some lemmas.\ fun lookup_pair :: "('a \ 'b) list \ 'a \ 'b::zero" where "lookup_pair [] x = 0"| "lookup_pair ((k, v) # xs) x = (case comp x k of Lt \ 0 | Eq \ v | Gt \ lookup_pair xs x)" fun update_by_pair :: "('a \ 'b) \ ('a \ 'b) list \ ('a \ 'b::zero) list" where "update_by_pair (k, v) [] = (if v = 0 then [] else [(k, v)])" | "update_by_pair (k, v) ((k', v') # xs) = (case comp k k' of Lt \ (if v = 0 then (k', v') # xs else (k, v) # (k', v') # xs) | Eq \ (if v = 0 then xs else (k, v) # xs) | Gt \ (k', v') # update_by_pair (k, v) xs)" (* TODO: Add update_by_gr_pair. *) definition sort_oalist :: "('a \ 'b) list \ ('a \ 'b::zero) list" where "sort_oalist xs = foldr update_by_pair xs []" fun update_by_fun_pair :: "'a \ ('b \ 'b) \ ('a \ 'b) list \ ('a \ 'b::zero) list" where "update_by_fun_pair k f [] = (let v = f 0 in if v = 0 then [] else [(k, v)])" | "update_by_fun_pair k f ((k', v') # xs) = (case comp k k' of Lt \ (let v = f 0 in if v = 0 then (k', v') # xs else (k, v) # (k', v') # xs) | Eq \ (let v = f v' in if v = 0 then xs else (k, v) # xs) | Gt \ (k', v') # update_by_fun_pair k f xs)" definition update_by_fun_gr_pair :: "'a \ ('b \ 'b) \ ('a \ 'b) list \ ('a \ 'b::zero) list" where "update_by_fun_gr_pair k f xs = (if xs = [] then (let v = f 0 in if v = 0 then [] else [(k, v)]) else if comp k (fst (last xs)) = Gt then (let v = f 0 in if v = 0 then xs else xs @ [(k, v)]) else update_by_fun_pair k f xs )" fun (in -) map_pair :: "(('a \ 'b) \ ('a \ 'c)) \ ('a \ 'b::zero) list \ ('a \ 'c::zero) list" where "map_pair f [] = []" | "map_pair f (kv # xs) = (let (k, v) = f kv; aux = map_pair f xs in if v = 0 then aux else (k, v) # aux)" text \The difference between @{const List.map} and @{const map_pair} is that the latter removes @{term 0} values, whereas the former does not.\ abbreviation (in -) map_val_pair :: "('a \ 'b \ 'c) \ ('a \ 'b::zero) list \ ('a \ 'c::zero) list" where "map_val_pair f \ map_pair (\(k, v). (k, f k v))" fun map2_val_pair :: "('a \ 'b \ 'c \ 'd) \ (('a \ 'b) list \ ('a \ 'd) list) \ (('a \ 'c) list \ ('a \ 'd) list) \ ('a \ 'b::zero) list \ ('a \ 'c::zero) list \ ('a \ 'd::zero) list" where "map2_val_pair f g h xs [] = g xs" | "map2_val_pair f g h [] ys = h ys" | "map2_val_pair f g h ((kx, vx) # xs) ((ky, vy) # ys) = (case comp kx ky of Lt \ (let v = f kx vx 0; aux = map2_val_pair f g h xs ((ky, vy) # ys) in if v = 0 then aux else (kx, v) # aux) | Eq \ (let v = f kx vx vy; aux = map2_val_pair f g h xs ys in if v = 0 then aux else (kx, v) # aux) | Gt \ (let v = f ky 0 vy; aux = map2_val_pair f g h ((kx, vx) # xs) ys in if v = 0 then aux else (ky, v) # aux))" fun lex_ord_pair :: "('a \ (('b, 'c) comp_opt)) \ (('a \ 'b::zero) list, ('a \ 'c::zero) list) comp_opt" where "lex_ord_pair f [] [] = Some Eq"| "lex_ord_pair f [] ((ky, vy) # ys) = (let aux = f ky 0 vy in if aux = Some Eq then lex_ord_pair f [] ys else aux)"| "lex_ord_pair f ((kx, vx) # xs) [] = (let aux = f kx vx 0 in if aux = Some Eq then lex_ord_pair f xs [] else aux)"| "lex_ord_pair f ((kx, vx) # xs) ((ky, vy) # ys) = (case comp kx ky of Lt \ (let aux = f kx vx 0 in if aux = Some Eq then lex_ord_pair f xs ((ky, vy) # ys) else aux) | Eq \ (let aux = f kx vx vy in if aux = Some Eq then lex_ord_pair f xs ys else aux) | Gt \ (let aux = f ky 0 vy in if aux = Some Eq then lex_ord_pair f ((kx, vx) # xs) ys else aux))" fun prod_ord_pair :: "('a \ 'b \ 'c \ bool) \ ('a \ 'b::zero) list \ ('a \ 'c::zero) list \ bool" where "prod_ord_pair f [] [] = True"| "prod_ord_pair f [] ((ky, vy) # ys) = (f ky 0 vy \ prod_ord_pair f [] ys)"| "prod_ord_pair f ((kx, vx) # xs) [] = (f kx vx 0 \ prod_ord_pair f xs [])"| "prod_ord_pair f ((kx, vx) # xs) ((ky, vy) # ys) = (case comp kx ky of Lt \ (f kx vx 0 \ prod_ord_pair f xs ((ky, vy) # ys)) | Eq \ (f kx vx vy \ prod_ord_pair f xs ys) | Gt \ (f ky 0 vy \ prod_ord_pair f ((kx, vx) # xs) ys))" text \@{const prod_ord_pair} is actually just a special case of @{const lex_ord_pair}, as proved below in lemma \prod_ord_pair_eq_lex_ord_pair\.\ subsubsection \@{const lookup_pair}\ lemma lookup_pair_eq_0: assumes "oalist_inv_raw xs" shows "lookup_pair xs k = 0 \ (k \ fst ` set xs)" using assms proof (induct xs rule: oalist_inv_raw_induct) case Nil show ?case by simp next case (Cons k' v' xs) show ?case proof (simp add: Cons(3) eq split: order.splits, rule, simp_all only: atomize_imp[symmetric]) assume "comp k k' = Lt" hence "k \ k'" by auto moreover have "k \ fst ` set xs" proof assume "k \ fst ` set xs" hence "lt k' k" by (rule Cons(4)) with \comp k k' = Lt\ show False by (simp add: Lt_lt_conv) qed ultimately show "k \ k' \ k \ fst ` set xs" .. next assume "comp k k' = Gt" hence "k \ k'" by auto thus "(lookup_pair xs k = 0) = (k \ k' \ k \ fst ` set xs)" by (simp add: Cons(5)) qed qed lemma lookup_pair_eq_value: assumes "oalist_inv_raw xs" and "v \ 0" shows "lookup_pair xs k = v \ ((k, v) \ set xs)" using assms(1) proof (induct xs rule: oalist_inv_raw_induct) case Nil from assms(2) show ?case by simp next case (Cons k' v' xs) have *: "(k', u) \ set xs" for u proof assume "(k', u) \ set xs" hence "fst (k', u) \ fst ` set xs" by fastforce hence "k' \ fst ` set xs" by simp hence "lt k' k'" by (rule Cons(4)) thus False by (simp add: lt_of_key_order_alt[symmetric]) qed show ?case proof (simp add: assms(2) Cons(5) eq split: order.split, intro conjI impI) assume "comp k k' = Lt" show "(k, v) \ set xs" proof assume "(k, v) \ set xs" hence "fst (k, v) \ fst ` set xs" by fastforce hence "k \ fst ` set xs" by simp hence "lt k' k" by (rule Cons(4)) with \comp k k' = Lt\ show False by (simp add: Lt_lt_conv) qed qed (auto simp: *) qed lemma lookup_pair_eq_valueI: assumes "oalist_inv_raw xs" and "(k, v) \ set xs" shows "lookup_pair xs k = v" proof - from assms(2) have "v \ snd ` set xs" by force moreover from assms(1) have "0 \ snd ` set xs" by (rule oalist_inv_rawD1) ultimately have "v \ 0" by blast with assms show ?thesis by (simp add: lookup_pair_eq_value) qed lemma lookup_dflt_eq_lookup_pair: assumes "oalist_inv_raw xs" shows "lookup_dflt xs = lookup_pair xs" proof (rule, simp add: lookup_dflt_def split: option.split, intro conjI impI allI) fix k assume "map_of xs k = None" with assms show "lookup_pair xs k = 0" by (simp add: lookup_pair_eq_0 map_of_eq_None_iff) next fix k v assume "map_of xs k = Some v" hence "(k, v) \ set xs" by (rule map_of_SomeD) with assms have "lookup_pair xs k = v" by (rule lookup_pair_eq_valueI) thus "v = lookup_pair xs k" by (rule HOL.sym) qed lemma lookup_pair_inj: assumes "oalist_inv_raw xs" and "oalist_inv_raw ys" and "lookup_pair xs = lookup_pair ys" shows "xs = ys" using assms proof (induct xs arbitrary: ys rule: oalist_inv_raw_induct) case Nil thus ?case proof (induct ys rule: oalist_inv_raw_induct) case Nil show ?case by simp next case (Cons k' v' ys) have "v' = lookup_pair ((k', v') # ys) k'" by simp also have "... = lookup_pair [] k'" by (simp only: Cons(6)) also have "... = 0" by simp finally have "v' = 0" . with Cons(3) show ?case .. qed next case *: (Cons k v xs) from *(6, 7) show ?case proof (induct ys rule: oalist_inv_raw_induct) case Nil have "v = lookup_pair ((k, v) # xs) k" by simp also have "... = lookup_pair [] k" by (simp only: Nil) also have "... = 0" by simp finally have "v = 0" . with *(3) show ?case .. next case (Cons k' v' ys) show ?case proof (cases "comp k k'") case Lt hence "\ lt k' k" by (simp add: Lt_lt_conv) with Cons(4) have "k \ fst ` set ys" by blast moreover from Lt have "k \ k'" by auto ultimately have "k \ fst ` set ((k', v') # ys)" by simp hence "0 = lookup_pair ((k', v') # ys) k" by (simp add: lookup_pair_eq_0[OF Cons(1)] del: lookup_pair.simps) also have "... = lookup_pair ((k, v) # xs) k" by (simp only: Cons(6)) also have "... = v" by simp finally have "v = 0" by simp with *(3) show ?thesis .. next case Eq hence "k' = k" by (simp only: eq) have "v' = lookup_pair ((k', v') # ys) k'" by simp also have "... = lookup_pair ((k, v) # xs) k" by (simp only: Cons(6) \k' = k\) also have "... = v" by simp finally have "v' = v" . moreover note \k' = k\ moreover from Cons(2) have "xs = ys" proof (rule *(5)) show "lookup_pair xs = lookup_pair ys" proof fix k0 show "lookup_pair xs k0 = lookup_pair ys k0" proof (cases "lt k k0") case True hence eq: "comp k0 k = Gt" by (simp add: Gt_lt_conv) have "lookup_pair xs k0 = lookup_pair ((k, v) # xs) k0" by (simp add: eq) also have "... = lookup_pair ((k, v') # ys) k0" by (simp only: Cons(6) \k' = k\) also have "... = lookup_pair ys k0" by (simp add: eq) finally show ?thesis . next case False with *(4) have "k0 \ fst ` set xs" by blast with *(2) have eq: "lookup_pair xs k0 = 0" by (simp add: lookup_pair_eq_0) from False Cons(4) have "k0 \ fst ` set ys" unfolding \k' = k\ by blast with Cons(2) have "lookup_pair ys k0 = 0" by (simp add: lookup_pair_eq_0) with eq show ?thesis by simp qed qed qed ultimately show ?thesis by simp next case Gt hence "\ lt k k'" by (simp add: Gt_lt_conv) with *(4) have "k' \ fst ` set xs" by blast moreover from Gt have "k' \ k" by auto ultimately have "k' \ fst ` set ((k, v) # xs)" by simp hence "0 = lookup_pair ((k, v) # xs) k'" by (simp add: lookup_pair_eq_0[OF *(1)] del: lookup_pair.simps) also have "... = lookup_pair ((k', v') # ys) k'" by (simp only: Cons(6)) also have "... = v'" by simp finally have "v' = 0" by simp with Cons(3) show ?thesis .. qed qed qed lemma lookup_pair_tl: assumes "oalist_inv_raw xs" shows "lookup_pair (tl xs) k = (if (\k'\fst ` set xs. le k k') then 0 else lookup_pair xs k)" proof - from assms have 1: "oalist_inv_raw (tl xs)" by (rule oalist_inv_raw_tl) show ?thesis proof (split if_split, intro conjI impI) assume *: "\x\fst ` set xs. le k x" show "lookup_pair (tl xs) k = 0" proof (simp add: lookup_pair_eq_0[OF 1], rule) assume k_in: "k \ fst ` set (tl xs)" hence "xs \ []" by auto then obtain k' v' ys where xs: "xs = (k', v') # ys" using prod.exhaust list.exhaust by metis have "k' \ fst ` set xs" unfolding xs by fastforce with * have "le k k'" .. from assms have "oalist_inv_raw ((k', v') # ys)" by (simp only: xs) moreover from k_in have "k \ fst ` set ys" by (simp add: xs) ultimately have "lt k' k" by (rule oalist_inv_raw_ConsD3) with \le k k'\ show False by simp qed next assume "\ (\k'\fst ` set xs. le k k')" hence "\x\fst ` set xs. \ le k x" by simp then obtain k'' where k''_in: "k'' \ fst ` set xs" and "\ le k k''" .. from this(2) have "lt k'' k" by simp from k''_in have "xs \ []" by auto then obtain k' v' ys where xs: "xs = (k', v') # ys" using prod.exhaust list.exhaust by metis from k''_in have "k'' = k' \ k'' \ fst ` set ys" by (simp add: xs) hence "lt k' k" proof assume "k'' = k'" with \lt k'' k\ show ?thesis by simp next from assms have "oalist_inv_raw ((k', v') # ys)" by (simp only: xs) moreover assume "k'' \ fst ` set ys" ultimately have "lt k' k''" by (rule oalist_inv_raw_ConsD3) thus ?thesis using \lt k'' k\ by (rule less_trans) qed hence "comp k k' = Gt" by (simp add: Gt_lt_conv) thus "lookup_pair (tl xs) k = lookup_pair xs k" by (simp add: xs lt_of_key_order_alt) qed qed lemma lookup_pair_tl': assumes "oalist_inv_raw xs" shows "lookup_pair (tl xs) k = (if k = fst (hd xs) then 0 else lookup_pair xs k)" proof - from assms have 1: "oalist_inv_raw (tl xs)" by (rule oalist_inv_raw_tl) show ?thesis proof (split if_split, intro conjI impI) assume k: "k = fst (hd xs)" show "lookup_pair (tl xs) k = 0" proof (simp add: lookup_pair_eq_0[OF 1], rule) assume k_in: "k \ fst ` set (tl xs)" hence "xs \ []" by auto then obtain k' v' ys where xs: "xs = (k', v') # ys" using prod.exhaust list.exhaust by metis from assms have "oalist_inv_raw ((k', v') # ys)" by (simp only: xs) moreover from k_in have "k' \ fst ` set ys" by (simp add: k xs) ultimately have "lt k' k'" by (rule oalist_inv_raw_ConsD3) thus False by simp qed next assume "k \ fst (hd xs)" show "lookup_pair (tl xs) k = lookup_pair xs k" proof (cases "xs = []") case True show ?thesis by (simp add: True) next case False then obtain k' v' ys where xs: "xs = (k', v') # ys" using prod.exhaust list.exhaust by metis show ?thesis proof (simp add: xs eq Lt_lt_conv split: order.split, intro conjI impI) from \k \ fst (hd xs)\ have "k \ k'" by (simp add: xs) moreover assume "k = k'" ultimately show "lookup_pair ys k' = v'" .. next assume "lt k k'" from assms have "oalist_inv_raw ys" unfolding xs by (rule oalist_inv_raw_ConsD1) moreover have "k \ fst ` set ys" proof assume "k \ fst ` set ys" with assms have "lt k' k" unfolding xs by (rule oalist_inv_raw_ConsD3) with \lt k k'\ show False by simp qed ultimately show "lookup_pair ys k = 0" by (simp add: lookup_pair_eq_0) qed qed qed qed lemma lookup_pair_filter: assumes "oalist_inv_raw xs" shows "lookup_pair (filter P xs) k = (let v = lookup_pair xs k in if P (k, v) then v else 0)" using assms proof (induct xs rule: oalist_inv_raw_induct) case Nil show ?case by simp next case (Cons k' v' xs) show ?case proof (simp add: Cons(5) Let_def eq split: order.split, intro conjI impI) show "lookup_pair xs k' = 0" proof (simp add: lookup_pair_eq_0 Cons(2), rule) assume "k' \ fst ` set xs" hence "lt k' k'" by (rule Cons(4)) thus False by simp qed next assume "comp k k' = Lt" hence "lt k k'" by (simp only: Lt_lt_conv) show "lookup_pair xs k = 0" proof (simp add: lookup_pair_eq_0 Cons(2), rule) assume "k \ fst ` set xs" hence "lt k' k" by (rule Cons(4)) with \lt k k'\ show False by simp qed qed qed lemma lookup_pair_map: assumes "oalist_inv_raw xs" and "\k'. snd (f (k', 0)) = 0" and "\a b. comp (fst (f a)) (fst (f b)) = comp (fst a) (fst b)" shows "lookup_pair (map f xs) (fst (f (k, v))) = snd (f (k, lookup_pair xs k))" using assms(1) proof (induct xs rule: oalist_inv_raw_induct) case Nil show ?case by (simp add: assms(2)) next case (Cons k' v' xs) obtain k'' v'' where f: "f (k', v') = (k'', v'')" by fastforce have "comp k k' = comp (fst (f (k, v))) (fst (f (k', v')))" by (simp add: assms(3)) also have "... = comp (fst (f (k, v))) k''" by (simp add: f) finally have eq0: "comp k k' = comp (fst (f (k, v))) k''" . show ?case proof (simp add: assms(2) split: order.split, intro conjI impI, simp add: eq) assume "k = k'" hence "lookup_pair (f (k', v') # map f xs) (fst (f (k', v))) = lookup_pair (f (k', v') # map f xs) (fst (f (k, v)))" by simp also have "... = snd (f (k', v'))" by (simp add: f eq0[symmetric], simp add: \k = k'\) finally show "lookup_pair (f (k', v') # map f xs) (fst (f (k', v))) = snd (f (k', v'))" . qed (simp_all add: f eq0 Cons(5)) qed lemma lookup_pair_Cons: assumes "oalist_inv_raw ((k, v) # xs)" shows "lookup_pair ((k, v) # xs) k0 = (if k = k0 then v else lookup_pair xs k0)" proof (simp add: eq split: order.split, intro impI) assume "comp k0 k = Lt" from assms have inv: "oalist_inv_raw xs" by (rule oalist_inv_raw_ConsD1) show "lookup_pair xs k0 = 0" proof (simp only: lookup_pair_eq_0[OF inv], rule) assume "k0 \ fst ` set xs" with assms have "lt k k0" by (rule oalist_inv_raw_ConsD3) with \comp k0 k = Lt\ show False by (simp add: Lt_lt_conv) qed qed lemma lookup_pair_single: "lookup_pair [(k, v)] k0 = (if k = k0 then v else 0)" by (simp add: eq split: order.split) subsubsection \@{const update_by_pair}\ lemma set_update_by_pair_subset: "set (update_by_pair kv xs) \ insert kv (set xs)" proof (induct xs arbitrary: kv) case Nil obtain k v where kv: "kv = (k, v)" by fastforce thus ?case by simp next case (Cons x xs) obtain k' v' where x: "x = (k', v')" by fastforce obtain k v where kv: "kv = (k, v)" by fastforce have 1: "set xs \ insert a (insert b (set xs))" for a b by auto have 2: "set (update_by_pair kv xs) \ insert kv (insert (k', v') (set xs))" for kv using Cons by blast show ?case by (simp add: x kv 1 2 split: order.split) qed lemma update_by_pair_sorted: assumes "sorted_wrt lt (map fst xs)" shows "sorted_wrt lt (map fst (update_by_pair kv xs))" using assms proof (induct xs arbitrary: kv) case Nil obtain k v where kv: "kv = (k, v)" by fastforce thus ?case by simp next case (Cons x xs) obtain k' v' where x: "x = (k', v')" by fastforce obtain k v where kv: "kv = (k, v)" by fastforce from Cons(2) have 1: "sorted_wrt lt (k' # (map fst xs))" by (simp add: x) hence 2: "sorted_wrt lt (map fst xs)" using sorted_wrt.elims(3) by fastforce hence 3: "sorted_wrt lt (map fst (update_by_pair (k, u) xs))" for u by (rule Cons(1)) have 4: "sorted_wrt lt (k' # map fst (update_by_pair (k, u) xs))" if *: "comp k k' = Gt" for u proof (simp, intro conjI ballI) fix y assume "y \ set (update_by_pair (k, u) xs)" also from set_update_by_pair_subset have "... \ insert (k, u) (set xs)" . finally have "y = (k, u) \ y \ set xs" by simp thus "lt k' (fst y)" proof assume "y = (k, u)" hence "fst y = k" by simp with * show ?thesis by (simp only: Gt_lt_conv) next from 1 have 5: "\y \ fst ` set xs. lt k' y" by simp assume "y \ set xs" hence "fst y \ fst ` set xs" by simp with 5 show ?thesis .. qed qed (fact 3) show ?case by (simp add: kv x 1 2 4 sorted_wrt2 split: order.split del: sorted_wrt.simps, intro conjI impI, simp add: 1 eq del: sorted_wrt.simps, simp add: Lt_lt_conv) qed lemma update_by_pair_not_0: assumes "0 \ snd ` set xs" shows "0 \ snd ` set (update_by_pair kv xs)" using assms proof (induct xs arbitrary: kv) case Nil obtain k v where kv: "kv = (k, v)" by fastforce thus ?case by simp next case (Cons x xs) obtain k' v' where x: "x = (k', v')" by fastforce obtain k v where kv: "kv = (k, v)" by fastforce from Cons(2) have 1: "v' \ 0" and 2: "0 \ snd ` set xs" by (auto simp: x) from 2 have 3: "0 \ snd ` set (update_by_pair (k, u) xs)" for u by (rule Cons(1)) show ?case by (auto simp: kv x 1 2 3 split: order.split) qed corollary oalist_inv_raw_update_by_pair: assumes "oalist_inv_raw xs" shows "oalist_inv_raw (update_by_pair kv xs)" proof (rule oalist_inv_rawI) from assms have "0 \ snd ` set xs" by (rule oalist_inv_rawD1) thus "0 \ snd ` set (update_by_pair kv xs)" by (rule update_by_pair_not_0) next from assms have "sorted_wrt lt (map fst xs)" by (rule oalist_inv_rawD2) thus "sorted_wrt lt (map fst (update_by_pair kv xs))" by (rule update_by_pair_sorted) qed lemma update_by_pair_less: assumes "v \ 0" and "xs = [] \ comp k (fst (hd xs)) = Lt" shows "update_by_pair (k, v) xs = (k, v) # xs" using assms(2) proof (induct xs) case Nil from assms(1) show ?case by simp next case (Cons x xs) obtain k' v' where x: "x = (k', v')" by fastforce from Cons(2) have "comp k k' = Lt" by (simp add: x) with assms(1) show ?case by (simp add: x) qed lemma lookup_pair_update_by_pair: assumes "oalist_inv_raw xs" shows "lookup_pair (update_by_pair (k1, v) xs) k2 = (if k1 = k2 then v else lookup_pair xs k2)" using assms proof (induct xs arbitrary: v rule: oalist_inv_raw_induct) case Nil show ?case by (simp split: order.split, simp add: eq) next case (Cons k' v' xs) show ?case proof (split if_split, intro conjI impI) assume "k1 = k2" with Cons(5) have eq0: "lookup_pair (update_by_pair (k2, u) xs) k2 = u" for u by (simp del: update_by_pair.simps) show "lookup_pair (update_by_pair (k1, v) ((k', v') # xs)) k2 = v" proof (simp add: \k1 = k2\ eq0 split: order.split, intro conjI impI) assume "comp k2 k' = Eq" hence "\ lt k' k2" by (simp add: eq) with Cons(4) have "k2 \ fst ` set xs" by auto thus "lookup_pair xs k2 = 0" using Cons(2) by (simp add: lookup_pair_eq_0) qed next assume "k1 \ k2" with Cons(5) have eq0: "lookup_pair (update_by_pair (k1, u) xs) k2 = lookup_pair xs k2" for u by (simp del: update_by_pair.simps) have *: "lookup_pair xs k2 = 0" if "lt k2 k'" proof - from \lt k2 k'\ have "\ lt k' k2" by auto with Cons(4) have "k2 \ fst ` set xs" by auto thus "lookup_pair xs k2 = 0" using Cons(2) by (simp add: lookup_pair_eq_0) qed show "lookup_pair (update_by_pair (k1, v) ((k', v') # xs)) k2 = lookup_pair ((k', v') # xs) k2" by (simp add: \k1 \ k2\ eq0 split: order.split, auto intro: * simp: \k1 \ k2\[symmetric] eq Gt_lt_conv Lt_lt_conv) qed qed corollary update_by_pair_id: assumes "oalist_inv_raw xs" and "lookup_pair xs k = v" shows "update_by_pair (k, v) xs = xs" proof (rule lookup_pair_inj, rule oalist_inv_raw_update_by_pair) show "lookup_pair (update_by_pair (k, v) xs) = lookup_pair xs" proof fix k0 from assms(2) show "lookup_pair (update_by_pair (k, v) xs) k0 = lookup_pair xs k0" by (auto simp: lookup_pair_update_by_pair[OF assms(1)]) qed qed fact+ lemma set_update_by_pair: assumes "oalist_inv_raw xs" and "v \ 0" shows "set (update_by_pair (k, v) xs) = insert (k, v) (set xs - range (Pair k))" (is "?A = ?B") proof (rule set_eqI) fix x::"'a \ 'b" obtain k' v' where x: "x = (k', v')" by fastforce from assms(1) have inv: "oalist_inv_raw (update_by_pair (k, v) xs)" by (rule oalist_inv_raw_update_by_pair) show "(x \ ?A) \ (x \ ?B)" proof (cases "v' = 0") case True have "0 \ snd ` set (update_by_pair (k, v) xs)" and "0 \ snd ` set xs" by (rule oalist_inv_rawD1, fact)+ hence "(k', 0) \ set (update_by_pair (k, v) xs)" and "(k', 0) \ set xs" using image_iff by fastforce+ thus ?thesis by (simp add: x True assms(2)) next case False show ?thesis by (auto simp: x lookup_pair_eq_value[OF inv False, symmetric] lookup_pair_eq_value[OF assms(1) False] lookup_pair_update_by_pair[OF assms(1)]) qed qed lemma set_update_by_pair_zero: assumes "oalist_inv_raw xs" shows "set (update_by_pair (k, 0) xs) = set xs - range (Pair k)" (is "?A = ?B") proof (rule set_eqI) fix x::"'a \ 'b" obtain k' v' where x: "x = (k', v')" by fastforce from assms(1) have inv: "oalist_inv_raw (update_by_pair (k, 0) xs)" by (rule oalist_inv_raw_update_by_pair) show "(x \ ?A) \ (x \ ?B)" proof (cases "v' = 0") case True have "0 \ snd ` set (update_by_pair (k, 0) xs)" and "0 \ snd ` set xs" by (rule oalist_inv_rawD1, fact)+ hence "(k', 0) \ set (update_by_pair (k, 0) xs)" and "(k', 0) \ set xs" using image_iff by fastforce+ thus ?thesis by (simp add: x True) next case False show ?thesis by (auto simp: x lookup_pair_eq_value[OF inv False, symmetric] lookup_pair_eq_value[OF assms False] lookup_pair_update_by_pair[OF assms] False) qed qed subsubsection \@{const update_by_fun_pair} and @{const update_by_fun_gr_pair}\ lemma update_by_fun_pair_eq_update_by_pair: assumes "oalist_inv_raw xs" shows "update_by_fun_pair k f xs = update_by_pair (k, f (lookup_pair xs k)) xs" using assms by (induct xs rule: oalist_inv_raw_induct, simp, simp split: order.split) corollary oalist_inv_raw_update_by_fun_pair: assumes "oalist_inv_raw xs" shows "oalist_inv_raw (update_by_fun_pair k f xs)" unfolding update_by_fun_pair_eq_update_by_pair[OF assms] using assms by (rule oalist_inv_raw_update_by_pair) corollary lookup_pair_update_by_fun_pair: assumes "oalist_inv_raw xs" shows "lookup_pair (update_by_fun_pair k1 f xs) k2 = (if k1 = k2 then f else id) (lookup_pair xs k2)" by (simp add: update_by_fun_pair_eq_update_by_pair[OF assms] lookup_pair_update_by_pair[OF assms]) lemma update_by_fun_pair_gr: assumes "oalist_inv_raw xs" and "xs = [] \ comp k (fst (last xs)) = Gt" shows "update_by_fun_pair k f xs = xs @ (if f 0 = 0 then [] else [(k, f 0)])" using assms proof (induct xs rule: oalist_inv_raw_induct) case Nil show ?case by simp next case (Cons k' v' xs) from Cons(6) have 1: "comp k (fst (last ((k', v') # xs))) = Gt" by simp have eq1: "comp k k' = Gt" proof (cases "xs = []") case True with 1 show ?thesis by simp next case False have "lt k' (fst (last xs))" by (rule Cons(4), simp add: False) from False 1 have "comp k (fst (last xs)) = Gt" by simp moreover from \lt k' (fst (last xs))\ have "comp (fst (last xs)) k' = Gt" by (simp add: Gt_lt_conv) ultimately show ?thesis by (meson Gt_lt_conv less_trans Lt_lt_conv[symmetric]) qed have eq2: "update_by_fun_pair k f xs = xs @ (if f 0 = 0 then [] else [(k, f 0)])" proof (rule Cons(5), simp only: disj_commute[of "xs = []"], rule disjCI) assume "xs \ []" with 1 show "comp k (fst (last xs)) = Gt" by simp qed show ?case by (simp split: order.split add: Let_def eq1 eq2) qed corollary update_by_fun_gr_pair_eq_update_by_fun_pair: assumes "oalist_inv_raw xs" shows "update_by_fun_gr_pair k f xs = update_by_fun_pair k f xs" by (simp add: update_by_fun_gr_pair_def Let_def update_by_fun_pair_gr[OF assms] split: order.split) corollary oalist_inv_raw_update_by_fun_gr_pair: assumes "oalist_inv_raw xs" shows "oalist_inv_raw (update_by_fun_gr_pair k f xs)" unfolding update_by_fun_pair_eq_update_by_pair[OF assms] update_by_fun_gr_pair_eq_update_by_fun_pair[OF assms] using assms by (rule oalist_inv_raw_update_by_pair) corollary lookup_pair_update_by_fun_gr_pair: assumes "oalist_inv_raw xs" shows "lookup_pair (update_by_fun_gr_pair k1 f xs) k2 = (if k1 = k2 then f else id) (lookup_pair xs k2)" by (simp add: update_by_fun_pair_eq_update_by_pair[OF assms] update_by_fun_gr_pair_eq_update_by_fun_pair[OF assms] lookup_pair_update_by_pair[OF assms]) subsubsection \@{const map_pair}\ lemma map_pair_cong: assumes "\kv. kv \ set xs \ f kv = g kv" shows "map_pair f xs = map_pair g xs" using assms proof (induct xs) case Nil show ?case by simp next case (Cons x xs) have "f x = g x" by (rule Cons(2), simp) moreover have "map_pair f xs = map_pair g xs" by (rule Cons(1), rule Cons(2), simp) ultimately show ?case by simp qed lemma map_pair_subset: "set (map_pair f xs) \ f ` set xs" proof (induct xs rule: map_pair.induct) case (1 f) show ?case by simp next case (2 f kv xs) obtain k v where f: "f kv = (k, v)" by fastforce from f[symmetric] HOL.refl have *: "set (map_pair f xs) \ f ` set xs" by (rule 2) show ?case by (simp add: f Let_def, intro conjI impI subset_insertI2 *) qed lemma oalist_inv_raw_map_pair: assumes "oalist_inv_raw xs" and "\a b. comp (fst (f a)) (fst (f b)) = comp (fst a) (fst b)" shows "oalist_inv_raw (map_pair f xs)" using assms(1) proof (induct xs rule: oalist_inv_raw_induct) case Nil from oalist_inv_raw_Nil show ?case by simp next case (Cons k v xs) obtain k' v' where f: "f (k, v) = (k', v')" by fastforce show ?case proof (simp add: f Let_def Cons(5), rule) assume "v' \ 0" with Cons(5) show "oalist_inv_raw ((k', v') # map_pair f xs)" proof (rule oalist_inv_raw_ConsI) assume "map_pair f xs \ []" hence "hd (map_pair f xs) \ set (map_pair f xs)" by simp also have "... \ f ` set xs" by (fact map_pair_subset) finally obtain x where "x \ set xs" and eq: "hd (map_pair f xs) = f x" .. from this(1) have "fst x \ fst ` set xs" by fastforce hence "lt k (fst x)" by (rule Cons(4)) hence "lt (fst (f (k, v))) (fst (f x))" by (simp add: Lt_lt_conv[symmetric] assms(2)) thus "lt k' (fst (hd (map_pair f xs)))" by (simp add: f eq) qed qed qed lemma lookup_pair_map_pair: assumes "oalist_inv_raw xs" and "snd (f (k, 0)) = 0" and "\a b. comp (fst (f a)) (fst (f b)) = comp (fst a) (fst b)" shows "lookup_pair (map_pair f xs) (fst (f (k, v))) = snd (f (k, lookup_pair xs k))" using assms(1) proof (induct xs rule: oalist_inv_raw_induct) case Nil show ?case by (simp add: assms(2)) next case (Cons k' v' xs) obtain k'' v'' where f: "f (k', v') = (k'', v'')" by fastforce have "comp (fst (f (k, v))) k'' = comp (fst (f (k, v))) (fst (f (k', v')))" by (simp add: f) also have "... = comp k k'" by (simp add: assms(3)) finally have eq0: "comp (fst (f (k, v))) k'' = comp k k'" . have *: "lookup_pair xs k = 0" if "comp k k' \ Gt" proof (simp add: lookup_pair_eq_0[OF Cons(2)], rule) assume "k \ fst ` set xs" hence "lt k' k" by (rule Cons(4)) hence "comp k k' = Gt" by (simp add: Gt_lt_conv) with \comp k k' \ Gt\ show False .. qed show ?case proof (simp add: assms(2) f Let_def eq0 Cons(5) split: order.split, intro conjI impI) assume "comp k k' = Lt" hence "comp k k' \ Gt" by simp hence "lookup_pair xs k = 0" by (rule *) thus "snd (f (k, lookup_pair xs k)) = 0" by (simp add: assms(2)) next assume "v'' = 0" assume "comp k k' = Eq" hence "k = k'" and "comp k k' \ Gt" by (simp only: eq, simp) from this(2) have "lookup_pair xs k = 0" by (rule *) hence "snd (f (k, lookup_pair xs k)) = 0" by (simp add: assms(2)) also have "... = snd (f (k, v'))" by (simp add: \k = k'\ f \v'' = 0\) finally show "snd (f (k, lookup_pair xs k)) = snd (f (k, v'))" . qed (simp add: f eq) qed lemma lookup_dflt_map_pair: assumes "distinct (map fst xs)" and "snd (f (k, 0)) = 0" and "\a b. (fst (f a) = fst (f b)) \ (fst a = fst b)" shows "lookup_dflt (map_pair f xs) (fst (f (k, v))) = snd (f (k, lookup_dflt xs k))" using assms(1) proof (induct xs) case Nil show ?case by (simp add: lookup_dflt_def assms(2)) next case (Cons x xs) obtain k' v' where x: "x = (k', v')" by fastforce obtain k'' v'' where f: "f (k', v') = (k'', v'')" by fastforce from Cons(2) have "distinct (map fst xs)" and "k' \ fst ` set xs" by (simp_all add: x) from this(1) have eq1: "lookup_dflt (map_pair f xs) (fst (f (k, v))) = snd (f (k, lookup_dflt xs k))" by (rule Cons(1)) have eq2: "lookup_dflt ((a, b) # ys) c = (if c = a then b else lookup_dflt ys c)" for a b c and ys::"('b \ 'e::zero) list" by (simp add: lookup_dflt_def map_of_Cons_code) from \k' \ fst ` set xs\ have "map_of xs k' = None" by (simp add: map_of_eq_None_iff) hence eq3: "lookup_dflt xs k' = 0" by (simp add: lookup_dflt_def) show ?case proof (simp add: f Let_def x eq1 eq2 eq3, intro conjI impI) assume "k = k'" hence "snd (f (k', 0)) = snd (f (k, 0))" by simp also have "... = 0" by (fact assms(2)) finally show "snd (f (k', 0)) = 0" . next assume "fst (f (k', v)) \ k''" hence "fst (f (k', v)) \ fst (f (k', v'))" by (simp add: f) thus "snd (f (k', 0)) = v''" by (simp add: assms(3)) next assume "k \ k'" assume "fst (f (k, v)) = k''" also have "... = fst (f (k', v'))" by (simp add: f) finally have "k = k'" by (simp add: assms(3)) with \k \ k'\ show "v'' = snd (f (k, lookup_dflt xs k))" .. qed qed lemma distinct_map_pair: assumes "distinct (map fst xs)" and "\a b. fst (f a) = fst (f b) \ fst a = fst b" shows "distinct (map fst (map_pair f xs))" using assms(1) proof (induct xs) case Nil show ?case by simp next case (Cons x xs) obtain k v where x: "x = (k, v)" by fastforce obtain k' v' where f: "f (k, v) = (k', v')" by fastforce from Cons(2) have "distinct (map fst xs)" and "k \ fst ` set xs" by (simp_all add: x) from this(1) have 1: "distinct (map fst (map_pair f xs))" by (rule Cons(1)) show ?case proof (simp add: x f Let_def 1, intro impI notI) assume "v' \ 0" assume "k' \ fst ` set (map_pair f xs)" then obtain y where "y \ set (map_pair f xs)" and "k' = fst y" .. from this(1) map_pair_subset have "y \ f ` set xs" .. then obtain z where "z \ set xs" and "y = f z" .. from this(2) have "fst (f z) = k'" by (simp add: \k' = fst y\) also have "... = fst (f (k, v))" by (simp add: f) finally have "fst z = fst (k, v)" by (rule assms(2)) also have "... = k" by simp finally have "k \ fst ` set xs" using \z \ set xs\ by blast with \k \ fst ` set xs\ show False .. qed qed lemma map_val_pair_cong: assumes "\k v. (k, v) \ set xs \ f k v = g k v" shows "map_val_pair f xs = map_val_pair g xs" proof (rule map_pair_cong) fix kv assume "kv \ set xs" moreover obtain k v where "kv = (k, v)" by fastforce ultimately show "(case kv of (k, v) \ (k, f k v)) = (case kv of (k, v) \ (k, g k v))" by (simp add: assms) qed lemma oalist_inv_raw_map_val_pair: assumes "oalist_inv_raw xs" shows "oalist_inv_raw (map_val_pair f xs)" by (rule oalist_inv_raw_map_pair, fact assms, auto) lemma lookup_pair_map_val_pair: assumes "oalist_inv_raw xs" and "f k 0 = 0" shows "lookup_pair (map_val_pair f xs) k = f k (lookup_pair xs k)" proof - let ?f = "\(k', v'). (k', f k' v')" have "lookup_pair (map_val_pair f xs) k = lookup_pair (map_val_pair f xs) (fst (?f (k, 0)))" by simp also have "... = snd (?f (k, local.lookup_pair xs k))" by (rule lookup_pair_map_pair, fact assms(1), auto simp: assms(2)) also have "... = f k (lookup_pair xs k)" by simp finally show ?thesis . qed lemma map_pair_id: assumes "oalist_inv_raw xs" shows "map_pair id xs = xs" using assms proof (induct xs rule: oalist_inv_raw_induct) case Nil show ?case by simp next case (Cons k v xs') show ?case by (simp add: Let_def Cons(3, 5) id_def[symmetric]) qed subsubsection \@{const map2_val_pair}\ definition map2_val_compat :: "(('a \ 'b::zero) list \ ('a \ 'c::zero) list) \ bool" where "map2_val_compat f \ (\zs. (oalist_inv_raw zs \ (oalist_inv_raw (f zs) \ fst ` set (f zs) \ fst ` set zs)))" lemma map2_val_compatI: assumes "\zs. oalist_inv_raw zs \ oalist_inv_raw (f zs)" and "\zs. oalist_inv_raw zs \ fst ` set (f zs) \ fst ` set zs" shows "map2_val_compat f" unfolding map2_val_compat_def using assms by blast lemma map2_val_compatD1: assumes "map2_val_compat f" and "oalist_inv_raw zs" shows "oalist_inv_raw (f zs)" using assms unfolding map2_val_compat_def by blast lemma map2_val_compatD2: assumes "map2_val_compat f" and "oalist_inv_raw zs" shows "fst ` set (f zs) \ fst ` set zs" using assms unfolding map2_val_compat_def by blast lemma map2_val_compat_Nil: assumes "map2_val_compat (f::('a \ 'b::zero) list \ ('a \ 'c::zero) list)" shows "f [] = []" proof - from assms oalist_inv_raw_Nil have "fst ` set (f []) \ fst ` set ([]::('a \ 'b) list)" by (rule map2_val_compatD2) thus ?thesis by simp qed lemma map2_val_compat_id: "map2_val_compat id" by (rule map2_val_compatI, auto) lemma map2_val_compat_map_val_pair: "map2_val_compat (map_val_pair f)" proof (rule map2_val_compatI, erule oalist_inv_raw_map_val_pair) fix zs from map_pair_subset image_iff show "fst ` set (map_val_pair f zs) \ fst ` set zs" by fastforce qed lemma fst_map2_val_pair_subset: assumes "oalist_inv_raw xs" and "oalist_inv_raw ys" assumes "map2_val_compat g" and "map2_val_compat h" shows "fst ` set (map2_val_pair f g h xs ys) \ fst ` set xs \ fst ` set ys" using assms proof (induct f g h xs ys rule: map2_val_pair.induct) case (1 f g h xs) show ?case by (simp, rule map2_val_compatD2, fact+) next case (2 f g h v va) show ?case by (simp del: set_simps(2), rule map2_val_compatD2, fact+) next case (3 f g h kx vx xs ky vy ys) from 3(4) have "oalist_inv_raw xs" by (rule oalist_inv_raw_ConsD1) from 3(5) have "oalist_inv_raw ys" by (rule oalist_inv_raw_ConsD1) show ?case proof (simp split: order.split, intro conjI impI) assume "comp kx ky = Lt" hence "fst ` set (map2_val_pair f g h xs ((ky, vy) # ys)) \ fst ` set xs \ fst ` set ((ky, vy) # ys)" using HOL.refl \oalist_inv_raw xs\ 3(5, 6, 7) by (rule 3(2)) thus "fst ` set (let v = f kx vx 0; aux = map2_val_pair f g h xs ((ky, vy) # ys) in if v = 0 then aux else (kx, v) # aux) \ insert ky (insert kx (fst ` set xs \ fst ` set ys))" by (auto simp: Let_def) next assume "comp kx ky = Eq" hence "fst ` set (map2_val_pair f g h xs ys) \ fst ` set xs \ fst ` set ys" using HOL.refl \oalist_inv_raw xs\ \oalist_inv_raw ys\ 3(6, 7) by (rule 3(1)) thus "fst ` set (let v = f kx vx vy; aux = map2_val_pair f g h xs ys in if v = 0 then aux else (kx, v) # aux) \ insert ky (insert kx (fst ` set xs \ fst ` set ys))" by (auto simp: Let_def) next assume "comp kx ky = Gt" hence "fst ` set (map2_val_pair f g h ((kx, vx) # xs) ys) \ fst ` set ((kx, vx) # xs) \ fst ` set ys" using HOL.refl 3(4) \oalist_inv_raw ys\ 3(6, 7) by (rule 3(3)) thus "fst ` set (let v = f ky 0 vy; aux = map2_val_pair f g h ((kx, vx) # xs) ys in if v = 0 then aux else (ky, v) # aux) \ insert ky (insert kx (fst ` set xs \ fst ` set ys))" by (auto simp: Let_def) qed qed lemma oalist_inv_raw_map2_val_pair: assumes "oalist_inv_raw xs" and "oalist_inv_raw ys" assumes "map2_val_compat g" and "map2_val_compat h" shows "oalist_inv_raw (map2_val_pair f g h xs ys)" using assms(1, 2) proof (induct xs arbitrary: ys rule: oalist_inv_raw_induct) case Nil show ?case proof (cases ys) case Nil show ?thesis by (simp add: Nil, rule map2_val_compatD1, fact assms(3), fact oalist_inv_raw_Nil) next case (Cons y ys') show ?thesis by (simp add: Cons, rule map2_val_compatD1, fact assms(4), simp only: Cons[symmetric], fact Nil) qed next case *: (Cons k v xs) from *(6) show ?case proof (induct ys rule: oalist_inv_raw_induct) case Nil show ?case by (simp, rule map2_val_compatD1, fact assms(3), fact *(1)) next case (Cons k' v' ys) show ?case proof (simp split: order.split, intro conjI impI) assume "comp k k' = Lt" hence 0: "lt k k'" by (simp only: Lt_lt_conv) from Cons(1) have 1: "oalist_inv_raw (map2_val_pair f g h xs ((k', v') # ys))" by (rule *(5)) show "oalist_inv_raw (let v = f k v 0; aux = map2_val_pair f g h xs ((k', v') # ys) in if v = 0 then aux else (k, v) # aux)" proof (simp add: Let_def, intro conjI impI) assume "f k v 0 \ 0" with 1 show "oalist_inv_raw ((k, f k v 0) # map2_val_pair f g h xs ((k', v') # ys))" proof (rule oalist_inv_raw_ConsI) define k0 where "k0 = fst (hd (local.map2_val_pair f g h xs ((k', v') # ys)))" assume "map2_val_pair f g h xs ((k', v') # ys) \ []" hence "k0 \ fst ` set (map2_val_pair f g h xs ((k', v') # ys))" by (simp add: k0_def) also from *(2) Cons(1) assms(3, 4) have "... \ fst ` set xs \ fst ` set ((k', v') # ys)" by (rule fst_map2_val_pair_subset) finally have "k0 \ fst ` set xs \ k0 = k' \ k0 \ fst ` set ys" by auto thus "lt k k0" proof (elim disjE) assume "k0 = k'" with 0 show ?thesis by simp next assume "k0 \ fst ` set ys" hence "lt k' k0" by (rule Cons(4)) with 0 show ?thesis by (rule less_trans) qed (rule *(4)) qed qed (rule 1) next assume "comp k k' = Eq" hence "k = k'" by (simp only: eq) from Cons(2) have 1: "oalist_inv_raw (map2_val_pair f g h xs ys)" by (rule *(5)) show "oalist_inv_raw (let v = f k v v'; aux = map2_val_pair f g h xs ys in if v = 0 then aux else (k, v) # aux)" proof (simp add: Let_def, intro conjI impI) assume "f k v v' \ 0" with 1 show "oalist_inv_raw ((k, f k v v') # map2_val_pair f g h xs ys)" proof (rule oalist_inv_raw_ConsI) define k0 where "k0 = fst (hd (map2_val_pair f g h xs ys))" assume "map2_val_pair f g h xs ys \ []" hence "k0 \ fst ` set (map2_val_pair f g h xs ys)" by (simp add: k0_def) also from *(2) Cons(2) assms(3, 4) have "... \ fst ` set xs \ fst ` set ys" by (rule fst_map2_val_pair_subset) finally show "lt k k0" proof assume "k0 \ fst ` set ys" hence "lt k' k0" by (rule Cons(4)) thus ?thesis by (simp only: \k = k'\) qed (rule *(4)) qed qed (rule 1) next assume "comp k k' = Gt" hence 0: "lt k' k" by (simp only: Gt_lt_conv) show "oalist_inv_raw (let va = f k' 0 v'; aux = map2_val_pair f g h ((k, v) # xs) ys in if va = 0 then aux else (k', va) # aux)" proof (simp add: Let_def, intro conjI impI) assume "f k' 0 v' \ 0" with Cons(5) show "oalist_inv_raw ((k', f k' 0 v') # map2_val_pair f g h ((k, v) # xs) ys)" proof (rule oalist_inv_raw_ConsI) define k0 where "k0 = fst (hd (map2_val_pair f g h ((k, v) # xs) ys))" assume "map2_val_pair f g h ((k, v) # xs) ys \ []" hence "k0 \ fst ` set (map2_val_pair f g h ((k, v) # xs) ys)" by (simp add: k0_def) also from *(1) Cons(2) assms(3, 4) have "... \ fst ` set ((k, v) # xs) \ fst ` set ys" by (rule fst_map2_val_pair_subset) finally have "k0 = k \ k0 \ fst ` set xs \ k0 \ fst ` set ys" by auto thus "lt k' k0" proof (elim disjE) assume "k0 = k" with 0 show ?thesis by simp next assume "k0 \ fst ` set xs" hence "lt k k0" by (rule *(4)) with 0 show ?thesis by (rule less_trans) qed (rule Cons(4)) qed qed (rule Cons(5)) qed qed qed lemma lookup_pair_map2_val_pair: assumes "oalist_inv_raw xs" and "oalist_inv_raw ys" assumes "map2_val_compat g" and "map2_val_compat h" assumes "\zs. oalist_inv_raw zs \ g zs = map_val_pair (\k v. f k v 0) zs" and "\zs. oalist_inv_raw zs \ h zs = map_val_pair (\k. f k 0) zs" and "\k. f k 0 0 = 0" shows "lookup_pair (map2_val_pair f g h xs ys) k0 = f k0 (lookup_pair xs k0) (lookup_pair ys k0)" using assms(1, 2) proof (induct xs arbitrary: ys rule: oalist_inv_raw_induct) case Nil show ?case proof (cases ys) case Nil show ?thesis by (simp add: Nil map2_val_compat_Nil[OF assms(3)] assms(7)) next case (Cons y ys') then obtain k v ys' where ys: "ys = (k, v) # ys'" by fastforce from Nil have "lookup_pair (h ys) k0 = lookup_pair (map_val_pair (\k. f k 0) ys) k0" by (simp only: assms(6)) also have "... = f k0 0 (lookup_pair ys k0)" by (rule lookup_pair_map_val_pair, fact Nil, fact assms(7)) finally have "lookup_pair (h ((k, v) # ys')) k0 = f k0 0 (lookup_pair ((k, v) # ys') k0)" by (simp only: ys) thus ?thesis by (simp add: ys) qed next case *: (Cons k v xs) from *(6) show ?case proof (induct ys rule: oalist_inv_raw_induct) case Nil from *(1) have "lookup_pair (g ((k, v) # xs)) k0 = lookup_pair (map_val_pair (\k v. f k v 0) ((k, v) # xs)) k0" by (simp only: assms(5)) also have "... = f k0 (lookup_pair ((k, v) # xs) k0) 0" by (rule lookup_pair_map_val_pair, fact *(1), fact assms(7)) finally show ?case by simp next case (Cons k' v' ys) show ?case proof (cases "comp k0 k = Lt \ comp k0 k' = Lt") case True hence 1: "comp k0 k = Lt" and 2: "comp k0 k' = Lt" by simp_all hence eq: "f k0 (lookup_pair ((k, v) # xs) k0) (lookup_pair ((k', v') # ys) k0) = 0" by (simp add: assms(7)) from *(1) Cons(1) assms(3, 4) have inv: "oalist_inv_raw (map2_val_pair f g h ((k, v) # xs) ((k', v') # ys))" by (rule oalist_inv_raw_map2_val_pair) show ?thesis proof (simp only: eq lookup_pair_eq_0[OF inv], rule) assume "k0 \ fst ` set (local.map2_val_pair f g h ((k, v) # xs) ((k', v') # ys))" also from *(1) Cons(1) assms(3, 4) have "... \ fst ` set ((k, v) # xs) \ fst ` set ((k', v') # ys)" by (rule fst_map2_val_pair_subset) finally have "k0 \ fst ` set xs \ k0 \ fst ` set ys" using 1 2 by auto thus False proof assume "k0 \ fst ` set xs" hence "lt k k0" by (rule *(4)) with 1 show ?thesis by (simp add: Lt_lt_conv) next assume "k0 \ fst ` set ys" hence "lt k' k0" by (rule Cons(4)) with 2 show ?thesis by (simp add: Lt_lt_conv) qed qed next case False show ?thesis proof (simp split: order.split del: lookup_pair.simps, intro conjI impI) assume "comp k k' = Lt" with False have "comp k0 k \ Lt" by (auto simp: Lt_lt_conv) show "lookup_pair (let v = f k v 0; aux = map2_val_pair f g h xs ((k', v') # ys) in if v = 0 then aux else (k, v) # aux) k0 = f k0 (lookup_pair ((k, v) # xs) k0) (lookup_pair ((k', v') # ys) k0)" proof (cases "comp k0 k") case Lt with \comp k0 k \ Lt\ show ?thesis .. next case Eq hence "k0 = k" by (simp only: eq) with \comp k k' = Lt\ have "comp k0 k' = Lt" by simp hence eq1: "lookup_pair ((k', v') # ys) k = 0" by (simp add: \k0 = k\) have eq2: "lookup_pair ((k, v) # xs) k = v" by simp show ?thesis proof (simp add: Let_def eq1 eq2 \k0 = k\ del: lookup_pair.simps, intro conjI impI) from *(2) Cons(1) assms(3, 4) have inv: "oalist_inv_raw (map2_val_pair f g h xs ((k', v') # ys))" by (rule oalist_inv_raw_map2_val_pair) show "lookup_pair (map2_val_pair f g h xs ((k', v') # ys)) k = 0" proof (simp only: lookup_pair_eq_0[OF inv], rule) assume "k \ fst ` set (local.map2_val_pair f g h xs ((k', v') # ys))" also from *(2) Cons(1) assms(3, 4) have "... \ fst ` set xs \ fst ` set ((k', v') # ys)" by (rule fst_map2_val_pair_subset) finally have "k \ fst ` set xs \ k \ fst ` set ys" using \comp k k' = Lt\ by auto thus False proof assume "k \ fst ` set xs" hence "lt k k" by (rule *(4)) thus ?thesis by simp next assume "k \ fst ` set ys" hence "lt k' k" by (rule Cons(4)) with \comp k k' = Lt\ show ?thesis by (simp add: Lt_lt_conv) qed qed qed simp next case Gt hence eq1: "lookup_pair ((k, v) # xs) k0 = lookup_pair xs k0" and eq2: "lookup_pair ((k, f k v 0) # map2_val_pair f g h xs ((k', v') # ys)) k0 = lookup_pair (map2_val_pair f g h xs ((k', v') # ys)) k0" by simp_all show ?thesis by (simp add: Let_def eq1 eq2 del: lookup_pair.simps, rule *(5), fact Cons(1)) qed next assume "comp k k' = Eq" hence "k = k'" by (simp only: eq) with False have "comp k0 k' \ Lt" by (auto simp: Lt_lt_conv) show "lookup_pair (let v = f k v v'; aux = map2_val_pair f g h xs ys in if v = 0 then aux else (k, v) # aux) k0 = f k0 (lookup_pair ((k, v) # xs) k0) (lookup_pair ((k', v') # ys) k0)" proof (cases "comp k0 k'") case Lt with \comp k0 k' \ Lt\ show ?thesis .. next case Eq hence "k0 = k'" by (simp only: eq) show ?thesis proof (simp add: Let_def \k = k'\ \k0 = k'\, intro impI) from *(2) Cons(2) assms(3, 4) have inv: "oalist_inv_raw (map2_val_pair f g h xs ys)" by (rule oalist_inv_raw_map2_val_pair) show "lookup_pair (map2_val_pair f g h xs ys) k' = 0" proof (simp only: lookup_pair_eq_0[OF inv], rule) assume "k' \ fst ` set (map2_val_pair f g h xs ys)" also from *(2) Cons(2) assms(3, 4) have "... \ fst ` set xs \ fst ` set ys" by (rule fst_map2_val_pair_subset) finally show False proof assume "k' \ fst ` set ys" hence "lt k' k'" by (rule Cons(4)) thus ?thesis by simp next assume "k' \ fst ` set xs" hence "lt k k'" by (rule *(4)) thus ?thesis by (simp add: \k = k'\) qed qed qed next case Gt hence eq1: "lookup_pair ((k, v) # xs) k0 = lookup_pair xs k0" and eq2: "lookup_pair ((k', v') # ys) k0 = lookup_pair ys k0" and eq3: "lookup_pair ((k, f k v v') # map2_val_pair f g h xs ys) k0 = lookup_pair (map2_val_pair f g h xs ys) k0" by (simp_all add: \k = k'\) show ?thesis by (simp add: Let_def eq1 eq2 eq3 del: lookup_pair.simps, rule *(5), fact Cons(2)) qed next assume "comp k k' = Gt" hence "comp k' k = Lt" by (simp only: Gt_lt_conv Lt_lt_conv) with False have "comp k0 k' \ Lt" by (auto simp: Lt_lt_conv) show "lookup_pair (let va = f k' 0 v'; aux = map2_val_pair f g h ((k, v) # xs) ys in if va = 0 then aux else (k', va) # aux) k0 = f k0 (lookup_pair ((k, v) # xs) k0) (lookup_pair ((k', v') # ys) k0)" proof (cases "comp k0 k'") case Lt with \comp k0 k' \ Lt\ show ?thesis .. next case Eq hence "k0 = k'" by (simp only: eq) with \comp k' k = Lt\ have "comp k0 k = Lt" by simp hence eq1: "lookup_pair ((k, v) # xs) k' = 0" by (simp add: \k0 = k'\) have eq2: "lookup_pair ((k', v') # ys) k' = v'" by simp show ?thesis proof (simp add: Let_def eq1 eq2 \k0 = k'\ del: lookup_pair.simps, intro conjI impI) from *(1) Cons(2) assms(3, 4) have inv: "oalist_inv_raw (map2_val_pair f g h ((k, v) # xs) ys)" by (rule oalist_inv_raw_map2_val_pair) show "lookup_pair (map2_val_pair f g h ((k, v) # xs) ys) k' = 0" proof (simp only: lookup_pair_eq_0[OF inv], rule) assume "k' \ fst ` set (map2_val_pair f g h ((k, v) # xs) ys)" also from *(1) Cons(2) assms(3, 4) have "... \ fst ` set ((k, v) # xs) \ fst ` set ys" by (rule fst_map2_val_pair_subset) finally have "k' \ fst ` set xs \ k' \ fst ` set ys" using \comp k' k = Lt\ by auto thus False proof assume "k' \ fst ` set ys" hence "lt k' k'" by (rule Cons(4)) thus ?thesis by simp next assume "k' \ fst ` set xs" hence "lt k k'" by (rule *(4)) with \comp k' k = Lt\ show ?thesis by (simp add: Lt_lt_conv) qed qed qed simp next case Gt hence eq1: "lookup_pair ((k', v') # ys) k0 = lookup_pair ys k0" and eq2: "lookup_pair ((k', f k' 0 v') # map2_val_pair f g h ((k, v) # xs) ys) k0 = lookup_pair (map2_val_pair f g h ((k, v) # xs) ys) k0" by simp_all show ?thesis by (simp add: Let_def eq1 eq2 del: lookup_pair.simps, rule Cons(5)) qed qed qed qed qed lemma map2_val_pair_singleton_eq_update_by_fun_pair: assumes "oalist_inv_raw xs" assumes "\k x. f k x 0 = x" and "\zs. oalist_inv_raw zs \ g zs = zs" and "h [(k, v)] = map_val_pair (\k. f k 0) [(k, v)]" shows "map2_val_pair f g h xs [(k, v)] = update_by_fun_pair k (\x. f k x v) xs" using assms(1) proof (induct xs rule: oalist_inv_raw_induct) case Nil show ?case by (simp add: Let_def assms(4)) next case (Cons k' v' xs) show ?case proof (cases "comp k' k") case Lt hence gr: "comp k k' = Gt" by (simp only: Gt_lt_conv Lt_lt_conv) show ?thesis by (simp add: Lt gr Let_def assms(2) Cons(3, 5)) next case Eq hence eq1: "comp k k' = Eq" and eq2: "k = k'" by (simp_all only: eq) show ?thesis by (simp add: Eq eq1 eq2 Let_def assms(3)[OF Cons(2)]) next case Gt hence less: "comp k k' = Lt" by (simp only: Gt_lt_conv Lt_lt_conv) show ?thesis by (simp add: Gt less Let_def assms(3)[OF Cons(1)]) qed qed subsubsection \@{const lex_ord_pair}\ lemma lex_ord_pair_EqI: assumes "oalist_inv_raw xs" and "oalist_inv_raw ys" and "\k. k \ fst ` set xs \ fst ` set ys \ f k (lookup_pair xs k) (lookup_pair ys k) = Some Eq" shows "lex_ord_pair f xs ys = Some Eq" using assms proof (induct xs arbitrary: ys rule: oalist_inv_raw_induct) case Nil thus ?case proof (induct ys rule: oalist_inv_raw_induct) case Nil show ?case by simp next case (Cons k v ys) show ?case proof (simp add: Let_def, intro conjI impI, rule Cons(5)) fix k0 assume "k0 \ fst ` set [] \ fst ` set ys" hence "k0 \ fst ` set ys" by simp hence "lt k k0" by (rule Cons(4)) hence "f k0 (lookup_pair [] k0) (lookup_pair ys k0) = f k0 (lookup_pair [] k0) (lookup_pair ((k, v) # ys) k0)" by (auto simp add: lookup_pair_Cons[OF Cons(1)] simp del: lookup_pair.simps) also have "... = Some Eq" by (rule Cons(6), simp add: \k0 \ fst ` set ys\) finally show "f k0 (lookup_pair [] k0) (lookup_pair ys k0) = Some Eq" . next have "f k 0 v = f k (lookup_pair [] k) (lookup_pair ((k, v) # ys) k)" by simp also have "... = Some Eq" by (rule Cons(6), simp) finally show "f k 0 v = Some Eq" . qed qed next case *: (Cons k v xs) from *(6, 7) show ?case proof (induct ys rule: oalist_inv_raw_induct) case Nil show ?case proof (simp add: Let_def, intro conjI impI, rule *(5), rule oalist_inv_raw_Nil) fix k0 assume "k0 \ fst ` set xs \ fst ` set []" hence "k0 \ fst ` set xs" by simp hence "lt k k0" by (rule *(4)) hence "f k0 (lookup_pair xs k0) (lookup_pair [] k0) = f k0 (lookup_pair ((k, v) # xs) k0) (lookup_pair [] k0)" by (auto simp add: lookup_pair_Cons[OF *(1)] simp del: lookup_pair.simps) also have "... = Some Eq" by (rule Nil, simp add: \k0 \ fst ` set xs\) finally show "f k0 (lookup_pair xs k0) (lookup_pair [] k0) = Some Eq" . next have "f k v 0 = f k (lookup_pair ((k, v) # xs) k) (lookup_pair [] k)" by simp also have "... = Some Eq" by (rule Nil, simp) finally show "f k v 0 = Some Eq" . qed next case (Cons k' v' ys) show ?case proof (simp split: order.split, intro conjI impI) assume "comp k k' = Lt" show "(let aux = f k v 0 in if aux = Some Eq then lex_ord_pair f xs ((k', v') # ys) else aux) = Some Eq" proof (simp add: Let_def, intro conjI impI, rule *(5), rule Cons(1)) fix k0 assume k0_in: "k0 \ fst ` set xs \ fst ` set ((k', v') # ys)" hence "k0 \ fst ` set xs \ k0 = k' \ k0 \ fst ` set ys" by auto hence "k0 \ k" proof (elim disjE) assume "k0 \ fst ` set xs" hence "lt k k0" by (rule *(4)) thus ?thesis by simp next assume "k0 = k'" with \comp k k' = Lt\ show ?thesis by auto next assume "k0 \ fst ` set ys" hence "lt k' k0" by (rule Cons(4)) with \comp k k' = Lt\ show ?thesis by (simp add: Lt_lt_conv) qed hence "f k0 (lookup_pair xs k0) (lookup_pair ((k', v') # ys) k0) = f k0 (lookup_pair ((k, v) # xs) k0) (lookup_pair ((k', v') # ys) k0)" by (auto simp add: lookup_pair_Cons[OF *(1)] simp del: lookup_pair.simps) also have "... = Some Eq" by (rule Cons(6), rule rev_subsetD, fact k0_in, auto) finally show "f k0 (lookup_pair xs k0) (lookup_pair ((k', v') # ys) k0) = Some Eq" . next have "f k v 0 = f k (lookup_pair ((k, v) # xs) k) (lookup_pair ((k', v') # ys) k)" by (simp add: \comp k k' = Lt\) also have "... = Some Eq" by (rule Cons(6), simp) finally show "f k v 0 = Some Eq" . qed next assume "comp k k' = Eq" hence "k = k'" by (simp only: eq) show "(let aux = f k v v' in if aux = Some Eq then lex_ord_pair f xs ys else aux) = Some Eq" proof (simp add: Let_def, intro conjI impI, rule *(5), rule Cons(2)) fix k0 assume k0_in: "k0 \ fst ` set xs \ fst ` set ys" hence "k0 \ k'" proof assume "k0 \ fst ` set xs" hence "lt k k0" by (rule *(4)) thus ?thesis by (simp add: \k = k'\) next assume "k0 \ fst ` set ys" hence "lt k' k0" by (rule Cons(4)) thus ?thesis by simp qed hence "f k0 (lookup_pair xs k0) (lookup_pair ys k0) = f k0 (lookup_pair ((k, v) # xs) k0) (lookup_pair ((k', v') # ys) k0)" by (simp add: lookup_pair_Cons[OF *(1)] lookup_pair_Cons[OF Cons(1)] del: lookup_pair.simps, auto simp: \k = k'\) also have "... = Some Eq" by (rule Cons(6), rule rev_subsetD, fact k0_in, auto) finally show "f k0 (lookup_pair xs k0) (lookup_pair ys k0) = Some Eq" . next have "f k v v' = f k (lookup_pair ((k, v) # xs) k) (lookup_pair ((k', v') # ys) k)" by (simp add: \k = k'\) also have "... = Some Eq" by (rule Cons(6), simp) finally show "f k v v' = Some Eq" . qed next assume "comp k k' = Gt" hence "comp k' k = Lt" by (simp only: Gt_lt_conv Lt_lt_conv) show "(let aux = f k' 0 v' in if aux = Some Eq then lex_ord_pair f ((k, v) # xs) ys else aux) = Some Eq" proof (simp add: Let_def, intro conjI impI, rule Cons(5)) fix k0 assume k0_in: "k0 \ fst ` set ((k, v) # xs) \ fst ` set ys" hence "k0 \ fst ` set xs \ k0 = k \ k0 \ fst ` set ys" by auto hence "k0 \ k'" proof (elim disjE) assume "k0 \ fst ` set xs" hence "lt k k0" by (rule *(4)) with \comp k' k = Lt\ show ?thesis by (simp add: Lt_lt_conv) next assume "k0 = k" with \comp k' k = Lt\ show ?thesis by auto next assume "k0 \ fst ` set ys" hence "lt k' k0" by (rule Cons(4)) thus ?thesis by simp qed hence "f k0 (lookup_pair ((k, v) # xs) k0) (lookup_pair ys k0) = f k0 (lookup_pair ((k, v) # xs) k0) (lookup_pair ((k', v') # ys) k0)" by (auto simp add: lookup_pair_Cons[OF Cons(1)] simp del: lookup_pair.simps) also have "... = Some Eq" by (rule Cons(6), rule rev_subsetD, fact k0_in, auto) finally show "f k0 (lookup_pair ((k, v) # xs) k0) (lookup_pair ys k0) = Some Eq" . next have "f k' 0 v' = f k' (lookup_pair ((k, v) # xs) k') (lookup_pair ((k', v') # ys) k')" by (simp add: \comp k' k = Lt\) also have "... = Some Eq" by (rule Cons(6), simp) finally show "f k' 0 v' = Some Eq" . qed qed qed qed lemma lex_ord_pair_valI: assumes "oalist_inv_raw xs" and "oalist_inv_raw ys" and "aux \ Some Eq" assumes "k \ fst ` set xs \ fst ` set ys" and "aux = f k (lookup_pair xs k) (lookup_pair ys k)" and "\k'. k' \ fst ` set xs \ fst ` set ys \ lt k' k \ f k' (lookup_pair xs k') (lookup_pair ys k') = Some Eq" shows "lex_ord_pair f xs ys = aux" using assms(1, 2, 4, 5, 6) proof (induct xs arbitrary: ys rule: oalist_inv_raw_induct) case Nil thus ?case proof (induct ys rule: oalist_inv_raw_induct) case Nil from Nil(1) show ?case by simp next case (Cons k' v' ys) from Cons(6) have "k = k' \ k \ fst ` set ys" by simp thus ?case proof assume "k = k'" with Cons(7) have "f k' 0 v' = aux" by simp thus ?thesis by (simp add: Let_def \k = k'\ assms(3)) next assume "k \ fst `set ys" hence "lt k' k" by (rule Cons(4)) hence "comp k k' = Gt" by (simp add: Gt_lt_conv) hence eq1: "lookup_pair ((k', v') # ys) k = lookup_pair ys k" by simp have "f k' (lookup_pair [] k') (lookup_pair ((k', v') # ys) k') = Some Eq" by (rule Cons(8), simp, fact) hence eq2: "f k' 0 v' = Some Eq" by simp show ?thesis proof (simp add: Let_def eq2, rule Cons(5)) from \k \ fst `set ys\ show "k \ fst ` set [] \ fst ` set ys" by simp next show "aux = f k (lookup_pair [] k) (lookup_pair ys k)" by (simp only: Cons(7) eq1) next fix k0 assume "lt k0 k" assume "k0 \ fst ` set [] \ fst ` set ys" hence k0_in: "k0 \ fst ` set ys" by simp hence "lt k' k0" by (rule Cons(4)) hence "comp k0 k' = Gt" by (simp add: Gt_lt_conv) hence "f k0 (lookup_pair [] k0) (lookup_pair ys k0) = f k0 (lookup_pair [] k0) (lookup_pair ((k', v') # ys) k0)" by simp also have "... = Some Eq" by (rule Cons(8), simp add: k0_in, fact) finally show "f k0 (lookup_pair [] k0) (lookup_pair ys k0) = Some Eq" . qed qed qed next case *: (Cons k' v' xs) from *(6, 7, 8, 9) show ?case proof (induct ys rule: oalist_inv_raw_induct) case Nil from Nil(1) have "k = k' \ k \ fst ` set xs" by simp thus ?case proof assume "k = k'" with Nil(2) have "f k' v' 0 = aux" by simp thus ?thesis by (simp add: Let_def \k = k'\ assms(3)) next assume "k \ fst ` set xs" hence "lt k' k" by (rule *(4)) hence "comp k k' = Gt" by (simp add: Gt_lt_conv) hence eq1: "lookup_pair ((k', v') # xs) k = lookup_pair xs k" by simp have "f k' (lookup_pair ((k', v') # xs) k') (lookup_pair [] k') = Some Eq" by (rule Nil(3), simp, fact) hence eq2: "f k' v' 0 = Some Eq" by simp show ?thesis proof (simp add: Let_def eq2, rule *(5), fact oalist_inv_raw_Nil) from \k \ fst `set xs\ show "k \ fst ` set xs \ fst ` set []" by simp next show "aux = f k (lookup_pair xs k) (lookup_pair [] k)" by (simp only: Nil(2) eq1) next fix k0 assume "lt k0 k" assume "k0 \ fst ` set xs \ fst ` set []" hence k0_in: "k0 \ fst ` set xs" by simp hence "lt k' k0" by (rule *(4)) hence "comp k0 k' = Gt" by (simp add: Gt_lt_conv) hence "f k0 (lookup_pair xs k0) (lookup_pair [] k0) = f k0 (lookup_pair ((k', v') # xs) k0) (lookup_pair [] k0)" by simp also have "... = Some Eq" by (rule Nil(3), simp add: k0_in, fact) finally show "f k0 (lookup_pair xs k0) (lookup_pair [] k0) = Some Eq" . qed qed next case (Cons k'' v'' ys) have 0: thesis if 1: "lt k k'" and 2: "lt k k''" for thesis proof - from 1 have "k \ k'" by simp moreover from 2 have "k \ k''" by simp ultimately have "k \ fst ` set xs \ k \ fst ` set ys" using Cons(6) by simp thus ?thesis proof assume "k \ fst ` set xs" hence "lt k' k" by (rule *(4)) with 1 show ?thesis by simp next assume "k \ fst ` set ys" hence "lt k'' k" by (rule Cons(4)) with 2 show ?thesis by simp qed qed show ?case proof (simp split: order.split, intro conjI impI) assume Lt: "comp k' k'' = Lt" show "(let aux = f k' v' 0 in if aux = Some Eq then lex_ord_pair f xs ((k'', v'') # ys) else aux) = aux" proof (simp add: Let_def split: order.split, intro conjI impI) assume "f k' v' 0 = Some Eq" have "k \ k'" proof assume "k = k'" have "aux = f k v' 0" by (simp add: Cons(7) \k = k'\ Lt) with \f k' v' 0 = Some Eq\ assms(3) show False by (simp add: \k = k'\) qed from Cons(1) show "lex_ord_pair f xs ((k'', v'') # ys) = aux" proof (rule *(5)) from Cons(6) \k \ k'\ show "k \ fst ` set xs \ fst ` set ((k'', v'') # ys)" by simp next show "aux = f k (lookup_pair xs k) (lookup_pair ((k'', v'') # ys) k)" by (simp add: Cons(7) lookup_pair_Cons[OF *(1)] \k \ k'\[symmetric] del: lookup_pair.simps) next fix k0 assume "lt k0 k" assume k0_in: "k0 \ fst ` set xs \ fst ` set ((k'', v'') # ys)" also have "... \ fst ` set ((k', v') # xs) \ fst ` set ((k'', v'') # ys)" by fastforce finally have k0_in': "k0 \ fst ` set ((k', v') # xs) \ fst ` set ((k'', v'') # ys)" . have "k' \ k0" proof assume "k' = k0" with k0_in have "k' \ fst ` set xs \ fst ` set ((k'', v'') # ys)" by simp with Lt have "k' \ fst ` set xs \ k' \ fst ` set ys" by auto thus False proof assume "k' \ fst ` set xs" hence "lt k' k'" by (rule *(4)) thus ?thesis by simp next assume "k' \ fst ` set ys" hence "lt k'' k'" by (rule Cons(4)) with Lt show ?thesis by (simp add: Lt_lt_conv) qed qed hence "f k0 (lookup_pair xs k0) (lookup_pair ((k'', v'') # ys) k0) = f k0 (lookup_pair ((k', v') # xs) k0) (lookup_pair ((k'', v'') # ys) k0)" by (simp add: lookup_pair_Cons[OF *(1)] del: lookup_pair.simps) also from k0_in' \lt k0 k\ have "... = Some Eq" by (rule Cons(8)) finally show "f k0 (lookup_pair xs k0) (lookup_pair ((k'', v'') # ys) k0) = Some Eq" . qed next assume "f k' v' 0 \ Some Eq" have "\ lt k' k" proof have "k' \ fst ` set ((k', v') # xs) \ fst ` set ((k'', v'') # ys)" by simp moreover assume "lt k' k" ultimately have "f k' (lookup_pair ((k', v') # xs) k') (lookup_pair ((k'', v'') # ys) k') = Some Eq" by (rule Cons(8)) hence "f k' v' 0 = Some Eq" by (simp add: Lt) with \f k' v' 0 \ Some Eq\ show False .. qed moreover have "\ lt k k'" proof assume "lt k k'" moreover from this Lt have "lt k k''" by (simp add: Lt_lt_conv) ultimately show False by (rule 0) qed ultimately have "k = k'" by simp show "f k' v' 0 = aux" by (simp add: Cons(7) \k = k'\ Lt) qed next assume "comp k' k'' = Eq" hence "k' = k''" by (simp only: eq) show "(let aux = f k' v' v'' in if aux = Some Eq then lex_ord_pair f xs ys else aux) = aux" proof (simp add: Let_def \k' = k''\ split: order.split, intro conjI impI) assume "f k'' v' v'' = Some Eq" have "k \ k''" proof assume "k = k''" have "aux = f k v' v''" by (simp add: Cons(7) \k = k''\ \k' = k''\) with \f k'' v' v'' = Some Eq\ assms(3) show False by (simp add: \k = k''\) qed from Cons(2) show "lex_ord_pair f xs ys = aux" proof (rule *(5)) from Cons(6) \k \ k''\ show "k \ fst ` set xs \ fst ` set ys" by (simp add: \k' = k''\) next show "aux = f k (lookup_pair xs k) (lookup_pair ys k)" by (simp add: Cons(7) lookup_pair_Cons[OF *(1)] lookup_pair_Cons[OF Cons(1)] del: lookup_pair.simps, simp add: \k' = k''\ \k \ k''\[symmetric]) next fix k0 assume "lt k0 k" assume k0_in: "k0 \ fst ` set xs \ fst ` set ys" also have "... \ fst ` set ((k', v') # xs) \ fst ` set ((k'', v'') # ys)" by fastforce finally have k0_in': "k0 \ fst ` set ((k', v') # xs) \ fst ` set ((k'', v'') # ys)" . have "k'' \ k0" proof assume "k'' = k0" with k0_in have "k'' \ fst ` set xs \ fst ` set ys" by simp thus False proof assume "k'' \ fst ` set xs" hence "lt k' k''" by (rule *(4)) thus ?thesis by (simp add: \k' = k''\) next assume "k'' \ fst ` set ys" hence "lt k'' k''" by (rule Cons(4)) thus ?thesis by simp qed qed hence "f k0 (lookup_pair xs k0) (lookup_pair ys k0) = f k0 (lookup_pair ((k', v') # xs) k0) (lookup_pair ((k'', v'') # ys) k0)" by (simp add: lookup_pair_Cons[OF *(1)] lookup_pair_Cons[OF Cons(1)] del: lookup_pair.simps, simp add: \k' = k''\) also from k0_in' \lt k0 k\ have "... = Some Eq" by (rule Cons(8)) finally show "f k0 (lookup_pair xs k0) (lookup_pair ys k0) = Some Eq" . qed next assume "f k'' v' v'' \ Some Eq" have "\ lt k'' k" proof have "k'' \ fst ` set ((k', v') # xs) \ fst ` set ((k'', v'') # ys)" by simp moreover assume "lt k'' k" ultimately have "f k'' (lookup_pair ((k', v') # xs) k'') (lookup_pair ((k'', v'') # ys) k'') = Some Eq" by (rule Cons(8)) hence "f k'' v' v'' = Some Eq" by (simp add: \k' = k''\) with \f k'' v' v'' \ Some Eq\ show False .. qed moreover have "\ lt k k''" proof assume "lt k k''" hence "lt k k'" by (simp only: \k' = k''\) thus False using \lt k k''\ by (rule 0) qed ultimately have "k = k''" by simp show "f k'' v' v'' = aux" by (simp add: Cons(7) \k = k''\ \k' = k''\) qed next assume Gt: "comp k' k'' = Gt" hence Lt: "comp k'' k' = Lt" by (simp only: Gt_lt_conv Lt_lt_conv) show "(let aux = f k'' 0 v'' in if aux = Some Eq then lex_ord_pair f ((k', v') # xs) ys else aux) = aux" proof (simp add: Let_def split: order.split, intro conjI impI) assume "f k'' 0 v'' = Some Eq" have "k \ k''" proof assume "k = k''" have "aux = f k 0 v''" by (simp add: Cons(7) \k = k''\ Lt) with \f k'' 0 v'' = Some Eq\ assms(3) show False by (simp add: \k = k''\) qed show "lex_ord_pair f ((k', v') # xs) ys = aux" proof (rule Cons(5)) from Cons(6) \k \ k''\ show "k \ fst ` set ((k', v') # xs) \ fst ` set ys" by simp next show "aux = f k (lookup_pair ((k', v') # xs) k) (lookup_pair ys k)" by (simp add: Cons(7) lookup_pair_Cons[OF Cons(1)] \k \ k''\[symmetric] del: lookup_pair.simps) next fix k0 assume "lt k0 k" assume k0_in: "k0 \ fst ` set ((k', v') # xs) \ fst ` set ys" also have "... \ fst ` set ((k', v') # xs) \ fst ` set ((k'', v'') # ys)" by fastforce finally have k0_in': "k0 \ fst ` set ((k', v') # xs) \ fst ` set ((k'', v'') # ys)" . have "k'' \ k0" proof assume "k'' = k0" with k0_in have "k'' \ fst ` set ((k', v') # xs) \ fst ` set ys" by simp with Lt have "k'' \ fst ` set xs \ k'' \ fst ` set ys" by auto thus False proof assume "k'' \ fst ` set xs" hence "lt k' k''" by (rule *(4)) with Lt show ?thesis by (simp add: Lt_lt_conv) next assume "k'' \ fst ` set ys" hence "lt k'' k''" by (rule Cons(4)) thus ?thesis by simp qed qed hence "f k0 (lookup_pair ((k', v') # xs) k0) (lookup_pair ys k0) = f k0 (lookup_pair ((k', v') # xs) k0) (lookup_pair ((k'', v'') # ys) k0)" by (simp add: lookup_pair_Cons[OF Cons(1)] del: lookup_pair.simps) also from k0_in' \lt k0 k\ have "... = Some Eq" by (rule Cons(8)) finally show "f k0 (lookup_pair ((k', v') # xs) k0) (lookup_pair ys k0) = Some Eq" . qed next assume "f k'' 0 v'' \ Some Eq" have "\ lt k'' k" proof have "k'' \ fst ` set ((k', v') # xs) \ fst ` set ((k'', v'') # ys)" by simp moreover assume "lt k'' k" ultimately have "f k'' (lookup_pair ((k', v') # xs) k'') (lookup_pair ((k'', v'') # ys) k'') = Some Eq" by (rule Cons(8)) hence "f k'' 0 v'' = Some Eq" by (simp add: Lt) with \f k'' 0 v'' \ Some Eq\ show False .. qed moreover have "\ lt k k''" proof assume "lt k k''" with Lt have "lt k k'" by (simp add: Lt_lt_conv) thus False using \lt k k''\ by (rule 0) qed ultimately have "k = k''" by simp show "f k'' 0 v'' = aux" by (simp add: Cons(7) \k = k''\ Lt) qed qed qed qed lemma lex_ord_pair_EqD: assumes "oalist_inv_raw xs" and "oalist_inv_raw ys" and "lex_ord_pair f xs ys = Some Eq" and "k \ fst ` set xs \ fst ` set ys" shows "f k (lookup_pair xs k) (lookup_pair ys k) = Some Eq" proof (rule ccontr) let ?A = "(fst ` set xs \ fst ` set ys) \ {k. f k (lookup_pair xs k) (lookup_pair ys k) \ Some Eq}" define k0 where "k0 = Min ?A" have "finite ?A" by auto assume "f k (lookup_pair xs k) (lookup_pair ys k) \ Some Eq" with assms(4) have "k \ ?A" by simp hence "?A \ {}" by blast with \finite ?A\ have "k0 \ ?A" unfolding k0_def by (rule Min_in) hence k0_in: "k0 \ fst ` set xs \ fst ` set ys" and neq: "f k0 (lookup_pair xs k0) (lookup_pair ys k0) \ Some Eq" by simp_all have "le k0 k'" if "k' \ ?A" for k' unfolding k0_def using \finite ?A\ that by (rule Min_le) hence "f k' (lookup_pair xs k') (lookup_pair ys k') = Some Eq" if "k' \ fst ` set xs \ fst ` set ys" and "lt k' k0" for k' using that by fastforce with assms(1, 2) neq k0_in HOL.refl have "lex_ord_pair f xs ys = f k0 (lookup_pair xs k0) (lookup_pair ys k0)" by (rule lex_ord_pair_valI) with assms(3) neq show False by simp qed lemma lex_ord_pair_valE: assumes "oalist_inv_raw xs" and "oalist_inv_raw ys" and "lex_ord_pair f xs ys = aux" and "aux \ Some Eq" obtains k where "k \ fst ` set xs \ fst ` set ys" and "aux = f k (lookup_pair xs k) (lookup_pair ys k)" and "\k'. k' \ fst ` set xs \ fst ` set ys \ lt k' k \ f k' (lookup_pair xs k') (lookup_pair ys k') = Some Eq" proof - let ?A = "(fst ` set xs \ fst ` set ys) \ {k. f k (lookup_pair xs k) (lookup_pair ys k) \ Some Eq}" define k where "k = Min ?A" have "finite ?A" by auto have "\k \ fst ` set xs \ fst ` set ys. f k (lookup_pair xs k) (lookup_pair ys k) \ Some Eq" (is ?prop) proof (rule ccontr) assume "\ ?prop" hence "f k (lookup_pair xs k) (lookup_pair ys k) = Some Eq" if "k \ fst ` set xs \ fst ` set ys" for k using that by auto with assms(1, 2) have "lex_ord_pair f xs ys = Some Eq" by (rule lex_ord_pair_EqI) with assms(3, 4) show False by simp qed then obtain k0 where "k0 \ fst ` set xs \ fst ` set ys" and "f k0 (lookup_pair xs k0) (lookup_pair ys k0) \ Some Eq" .. hence "k0 \ ?A" by simp hence "?A \ {}" by blast with \finite ?A\ have "k \ ?A" unfolding k_def by (rule Min_in) hence k_in: "k \ fst ` set xs \ fst ` set ys" and neq: "f k (lookup_pair xs k) (lookup_pair ys k) \ Some Eq" by simp_all have "le k k'" if "k' \ ?A" for k' unfolding k_def using \finite ?A\ that by (rule Min_le) hence *: "\k'. k' \ fst ` set xs \ fst ` set ys \ lt k' k \ f k' (lookup_pair xs k') (lookup_pair ys k') = Some Eq" by fastforce with assms(1, 2) neq k_in HOL.refl have "lex_ord_pair f xs ys = f k (lookup_pair xs k) (lookup_pair ys k)" by (rule lex_ord_pair_valI) hence "aux = f k (lookup_pair xs k) (lookup_pair ys k)" by (simp only: assms(3)) with k_in show ?thesis using * .. qed subsubsection \@{const prod_ord_pair}\ lemma prod_ord_pair_eq_lex_ord_pair: "prod_ord_pair P xs ys = (lex_ord_pair (\k x y. if P k x y then Some Eq else None) xs ys = Some Eq)" proof (induct P xs ys rule: prod_ord_pair.induct) case (1 P) show ?case by simp next case (2 P ky vy ys) thus ?case by simp next case (3 P kx vx xs) thus ?case by simp next case (4 P kx vx xs ky vy ys) show ?case proof (cases "comp kx ky") case Lt thus ?thesis by (simp add: 4(2)[OF Lt]) next case Eq thus ?thesis by (simp add: 4(1)[OF Eq]) next case Gt thus ?thesis by (simp add: 4(3)[OF Gt]) qed qed lemma prod_ord_pairI: assumes "oalist_inv_raw xs" and "oalist_inv_raw ys" and "\k. k \ fst ` set xs \ fst ` set ys \ P k (lookup_pair xs k) (lookup_pair ys k)" shows "prod_ord_pair P xs ys" unfolding prod_ord_pair_eq_lex_ord_pair by (rule lex_ord_pair_EqI, fact, fact, simp add: assms(3)) lemma prod_ord_pairD: assumes "oalist_inv_raw xs" and "oalist_inv_raw ys" and "prod_ord_pair P xs ys" and "k \ fst ` set xs \ fst ` set ys" shows "P k (lookup_pair xs k) (lookup_pair ys k)" proof - from assms have "(if P k (lookup_pair xs k) (lookup_pair ys k) then Some Eq else None) = Some Eq" unfolding prod_ord_pair_eq_lex_ord_pair by (rule lex_ord_pair_EqD) thus ?thesis by (simp split: if_splits) qed corollary prod_ord_pair_alt: assumes "oalist_inv_raw xs" and "oalist_inv_raw ys" shows "(prod_ord_pair P xs ys) \ (\k\fst ` set xs \ fst ` set ys. P k (lookup_pair xs k) (lookup_pair ys k))" using prod_ord_pairI[OF assms] prod_ord_pairD[OF assms] by meson subsubsection \@{const sort_oalist}\ lemma oalist_inv_raw_foldr_update_by_pair: assumes "oalist_inv_raw ys" shows "oalist_inv_raw (foldr update_by_pair xs ys)" proof (induct xs) case Nil from assms show ?case by simp next case (Cons x xs) hence "oalist_inv_raw (update_by_pair x (foldr update_by_pair xs ys))" by (rule oalist_inv_raw_update_by_pair) thus ?case by simp qed corollary oalist_inv_raw_sort_oalist: "oalist_inv_raw (sort_oalist xs)" proof - from oalist_inv_raw_Nil have "oalist_inv_raw (foldr local.update_by_pair xs [])" by (rule oalist_inv_raw_foldr_update_by_pair) thus "oalist_inv_raw (sort_oalist xs)" by (simp only: sort_oalist_def) qed lemma sort_oalist_id: assumes "oalist_inv_raw xs" shows "sort_oalist xs = xs" proof - have "foldr update_by_pair xs ys = xs @ ys" if "oalist_inv_raw (xs @ ys)" for ys using assms that proof (induct xs rule: oalist_inv_raw_induct) case Nil show ?case by simp next case (Cons k v xs) from Cons(6) have *: "oalist_inv_raw ((k, v) # (xs @ ys))" by simp hence 1: "oalist_inv_raw (xs @ ys)" by (rule oalist_inv_raw_ConsD1) hence 2: "foldr update_by_pair xs ys = xs @ ys" by (rule Cons(5)) show ?case proof (simp add: 2, rule update_by_pair_less) from * show "v \ 0" by (auto simp: oalist_inv_raw_def) next have "comp k (fst (hd (xs @ ys))) = Lt \ xs @ ys = []" proof (rule disjCI) assume "xs @ ys \ []" then obtain k'' v'' zs where eq0: "xs @ ys = (k'', v'') # zs" using list.exhaust prod.exhaust by metis from * have "lt k k''" by (simp add: eq0 oalist_inv_raw_def) thus "comp k (fst (hd (xs @ ys))) = Lt" by (simp add: eq0 Lt_lt_conv) qed thus "xs @ ys = [] \ comp k (fst (hd (xs @ ys))) = Lt" by auto qed qed with assms show ?thesis by (simp add: sort_oalist_def) qed lemma set_sort_oalist: assumes "distinct (map fst xs)" shows "set (sort_oalist xs) = {kv. kv \ set xs \ snd kv \ 0}" using assms proof (induct xs) case Nil show ?case by (simp add: sort_oalist_def) next case (Cons x xs) obtain k v where x: "x = (k, v)" by fastforce from Cons(2) have "distinct (map fst xs)" and "k \ fst ` set xs" by (simp_all add: x) from this(1) have "set (sort_oalist xs) = {kv \ set xs. snd kv \ 0}" by (rule Cons(1)) with \k \ fst ` set xs\ have eq: "set (sort_oalist xs) - range (Pair k) = {kv \ set xs. snd kv \ 0}" by (auto simp: image_iff) have "set (sort_oalist (x # xs)) = set (update_by_pair (k, v) (sort_oalist xs))" by (simp add: sort_oalist_def x) also have "... = {kv \ set (x # xs). snd kv \ 0}" proof (cases "v = 0") case True have "set (update_by_pair (k, v) (sort_oalist xs)) = set (sort_oalist xs) - range (Pair k)" unfolding True using oalist_inv_raw_sort_oalist by (rule set_update_by_pair_zero) also have "... = {kv \ set (x # xs). snd kv \ 0}" by (auto simp: eq x True) finally show ?thesis . next case False with oalist_inv_raw_sort_oalist have "set (update_by_pair (k, v) (sort_oalist xs)) = insert (k, v) (set (sort_oalist xs) - range (Pair k))" by (rule set_update_by_pair) also have "... = {kv \ set (x # xs). snd kv \ 0}" by (auto simp: eq x False) finally show ?thesis . qed finally show ?case . qed lemma lookup_pair_sort_oalist': assumes "distinct (map fst xs)" shows "lookup_pair (sort_oalist xs) = lookup_dflt xs" using assms proof (induct xs) case Nil show ?case by (simp add: sort_oalist_def lookup_dflt_def) next case (Cons x xs) obtain k v where x: "x = (k, v)" by fastforce from Cons(2) have "distinct (map fst xs)" and "k \ fst ` set xs" by (simp_all add: x) from this(1) have eq1: "lookup_pair (sort_oalist xs) = lookup_dflt xs" by (rule Cons(1)) have eq2: "sort_oalist (x # xs) = update_by_pair (k, v) (sort_oalist xs)" by (simp add: x sort_oalist_def) show ?case proof fix k' have "lookup_pair (sort_oalist (x # xs)) k' = (if k = k' then v else lookup_dflt xs k')" by (simp add: eq1 eq2 lookup_pair_update_by_pair[OF oalist_inv_raw_sort_oalist]) also have "... = lookup_dflt (x # xs) k'" by (simp add: x lookup_dflt_def) finally show "lookup_pair (sort_oalist (x # xs)) k' = lookup_dflt (x # xs) k'" . qed qed end locale comparator2 = comparator comp1 + cmp2: comparator comp2 for comp1 comp2 :: "'a comparator" begin lemma set_sort_oalist: assumes "cmp2.oalist_inv_raw xs" shows "set (sort_oalist xs) = set xs" proof - have rl: "set (foldr update_by_pair xs ys) = set xs \ set ys" if "oalist_inv_raw ys" and "fst ` set xs \ fst ` set ys = {}" for ys using assms that(2) proof (induct xs rule: cmp2.oalist_inv_raw_induct) case Nil show ?case by simp next case (Cons k v xs) from Cons(6) have "k \ fst ` set ys" and "fst ` set xs \ fst ` set ys = {}" by simp_all from this(2) have eq1: "set (foldr update_by_pair xs ys) = set xs \ set ys" by (rule Cons(5)) have "\ cmp2.lt k k" by auto with Cons(4) have "k \ fst ` set xs" by blast with \k \ fst ` set ys\ have "k \ fst ` (set xs \ set ys)" by (simp add: image_Un) hence "(set xs \ set ys) \ range (Pair k) = {}" by (smt Int_emptyI fstI image_iff) hence eq2: "(set xs \ set ys) - range (Pair k) = set xs \ set ys" by (rule Diff_triv) from \oalist_inv_raw ys\ have "oalist_inv_raw (foldr update_by_pair xs ys)" by (rule oalist_inv_raw_foldr_update_by_pair) hence "set (update_by_pair (k, v) (foldr update_by_pair xs ys)) = insert (k, v) (set (foldr update_by_pair xs ys) - range (Pair k))" using Cons(3) by (rule set_update_by_pair) also have "... = insert (k, v) (set xs \ set ys)" by (simp only: eq1 eq2) finally show ?case by simp qed have "set (foldr update_by_pair xs []) = set xs \ set []" by (rule rl, fact oalist_inv_raw_Nil, simp) thus ?thesis by (simp add: sort_oalist_def) qed lemma lookup_pair_eqI: assumes "oalist_inv_raw xs" and "cmp2.oalist_inv_raw ys" and "set xs = set ys" shows "lookup_pair xs = cmp2.lookup_pair ys" proof fix k show "lookup_pair xs k = cmp2.lookup_pair ys k" proof (cases "cmp2.lookup_pair ys k = 0") case True with assms(2) have "k \ fst ` set ys" by (simp add: cmp2.lookup_pair_eq_0) with assms(1) show ?thesis by (simp add: True assms(3)[symmetric] lookup_pair_eq_0) next case False define v where "v = cmp2.lookup_pair ys k" from False have "v \ 0" by (simp add: v_def) with assms(2) v_def[symmetric] have "(k, v) \ set ys" by (simp add: cmp2.lookup_pair_eq_value) with assms(1) \v \ 0\ have "lookup_pair xs k = v" by (simp add: assms(3)[symmetric] lookup_pair_eq_value) thus ?thesis by (simp only: v_def) qed qed corollary lookup_pair_sort_oalist: assumes "cmp2.oalist_inv_raw xs" shows "lookup_pair (sort_oalist xs) = cmp2.lookup_pair xs" by (rule lookup_pair_eqI, rule oalist_inv_raw_sort_oalist, fact, rule set_sort_oalist, fact) end (* comparator2 *) subsection \Invariant on Pairs\ type_synonym ('a, 'b, 'c) oalist_raw = "('a \ 'b) list \ 'c" locale oalist_raw = fixes rep_key_order::"'o \ 'a key_order" begin sublocale comparator "key_compare (rep_key_order x)" by (fact comparator_key_compare) definition oalist_inv :: "('a, 'b::zero, 'o) oalist_raw \ bool" where "oalist_inv xs \ oalist_inv_raw (snd xs) (fst xs)" lemma oalist_inv_alt: "oalist_inv (xs, ko) \ oalist_inv_raw ko xs" by (simp add: oalist_inv_def) subsection \Operations on Raw Ordered Associative Lists\ fun sort_oalist_aux :: "'o \ ('a, 'b, 'o) oalist_raw \ ('a \ 'b::zero) list" where "sort_oalist_aux ko (xs, ox) = (if ko = ox then xs else sort_oalist ko xs)" fun lookup_raw :: "('a, 'b, 'o) oalist_raw \ 'a \ 'b::zero" where "lookup_raw (xs, ko) = lookup_pair ko xs" definition sorted_domain_raw :: "'o \ ('a, 'b::zero, 'o) oalist_raw \ 'a list" where "sorted_domain_raw ko xs = map fst (sort_oalist_aux ko xs)" fun tl_raw :: "('a, 'b, 'o) oalist_raw \ ('a, 'b::zero, 'o) oalist_raw" where "tl_raw (xs, ko) = (List.tl xs, ko)" fun min_key_val_raw :: "'o \ ('a, 'b, 'o) oalist_raw \ ('a \ 'b::zero)" where "min_key_val_raw ko (xs, ox) = (if ko = ox then List.hd else min_list_param (\x y. le ko (fst x) (fst y))) xs" fun update_by_raw :: "('a \ 'b) \ ('a, 'b, 'o) oalist_raw \ ('a, 'b::zero, 'o) oalist_raw" where "update_by_raw kv (xs, ko) = (update_by_pair ko kv xs, ko)" fun update_by_fun_raw :: "'a \ ('b \ 'b) \ ('a, 'b, 'o) oalist_raw \ ('a, 'b::zero, 'o) oalist_raw" where "update_by_fun_raw k f (xs, ko) = (update_by_fun_pair ko k f xs, ko)" fun update_by_fun_gr_raw :: "'a \ ('b \ 'b) \ ('a, 'b, 'o) oalist_raw \ ('a, 'b::zero, 'o) oalist_raw" where "update_by_fun_gr_raw k f (xs, ko) = (update_by_fun_gr_pair ko k f xs, ko)" fun (in -) filter_raw :: "('a \ bool) \ ('a list \ 'b) \ ('a list \ 'b)" where "filter_raw P (xs, ko) = (filter P xs, ko)" fun (in -) map_raw :: "(('a \ 'b) \ ('a \ 'c)) \ (('a \ 'b::zero) list \ 'd) \ ('a \ 'c::zero) list \ 'd" where "map_raw f (xs, ko) = (map_pair f xs, ko)" abbreviation (in -) "map_val_raw f \ map_raw (\(k, v). (k, f k v))" fun map2_val_raw :: "('a \ 'b \ 'c \ 'd) \ (('a, 'b, 'o) oalist_raw \ ('a, 'd, 'o) oalist_raw) \ (('a, 'c, 'o) oalist_raw \ ('a, 'd, 'o) oalist_raw) \ ('a, 'b::zero, 'o) oalist_raw \ ('a, 'c::zero, 'o) oalist_raw \ ('a, 'd::zero, 'o) oalist_raw" where "map2_val_raw f g h (xs, ox) ys = (map2_val_pair ox f (\zs. fst (g (zs, ox))) (\zs. fst (h (zs, ox))) xs (sort_oalist_aux ox ys), ox)" definition lex_ord_raw :: "'o \ ('a \ (('b, 'c) comp_opt)) \ (('a, 'b::zero, 'o) oalist_raw, ('a, 'c::zero, 'o) oalist_raw) comp_opt" where "lex_ord_raw ko f xs ys = lex_ord_pair ko f (sort_oalist_aux ko xs) (sort_oalist_aux ko ys)" fun prod_ord_raw :: "('a \ 'b \ 'c \ bool) \ ('a, 'b::zero, 'o) oalist_raw \ ('a, 'c::zero, 'o) oalist_raw \ bool" where "prod_ord_raw f (xs, ox) ys = prod_ord_pair ox f xs (sort_oalist_aux ox ys)" fun oalist_eq_raw :: "('a, 'b, 'o) oalist_raw \ ('a, 'b::zero, 'o) oalist_raw \ bool" where "oalist_eq_raw (xs, ox) ys = (xs = (sort_oalist_aux ox ys))" fun sort_oalist_raw :: "('a, 'b, 'o) oalist_raw \ ('a, 'b::zero, 'o) oalist_raw" where "sort_oalist_raw (xs, ko) = (sort_oalist ko xs, ko)" subsubsection \@{const sort_oalist_aux}\ lemma set_sort_oalist_aux: assumes "oalist_inv xs" shows "set (sort_oalist_aux ko xs) = set (fst xs)" proof - obtain xs' ko' where xs: "xs = (xs', ko')" by fastforce interpret ko2: comparator2 "key_compare (rep_key_order ko)" "key_compare (rep_key_order ko')" .. from assms show ?thesis by (simp add: xs oalist_inv_alt ko2.set_sort_oalist) qed lemma oalist_inv_raw_sort_oalist_aux: assumes "oalist_inv xs" shows "oalist_inv_raw ko (sort_oalist_aux ko xs)" proof - obtain xs' ko' where xs: "xs = (xs', ko')" by fastforce from assms show ?thesis by (simp add: xs oalist_inv_alt oalist_inv_raw_sort_oalist) qed lemma oalist_inv_sort_oalist_aux: assumes "oalist_inv xs" shows "oalist_inv (sort_oalist_aux ko xs, ko)" unfolding oalist_inv_alt using assms by (rule oalist_inv_raw_sort_oalist_aux) lemma lookup_pair_sort_oalist_aux: assumes "oalist_inv xs" shows "lookup_pair ko (sort_oalist_aux ko xs) = lookup_raw xs" proof - obtain xs' ko' where xs: "xs = (xs', ko')" by fastforce interpret ko2: comparator2 "key_compare (rep_key_order ko)" "key_compare (rep_key_order ko')" .. from assms show ?thesis by (simp add: xs oalist_inv_alt ko2.lookup_pair_sort_oalist) qed subsubsection \@{const lookup_raw}\ lemma lookup_raw_eq_value: assumes "oalist_inv xs" and "v \ 0" shows "lookup_raw xs k = v \ ((k, v) \ set (fst xs))" proof - obtain xs' ox where xs: "xs = (xs', ox)" by fastforce from assms(1) have "oalist_inv_raw ox xs'" by (simp add: xs oalist_inv_def) show ?thesis by (simp add: xs, rule lookup_pair_eq_value, fact+) qed lemma lookup_raw_eq_valueI: assumes "oalist_inv xs" and "(k, v) \ set (fst xs)" shows "lookup_raw xs k = v" proof - obtain xs' ox where xs: "xs = (xs', ox)" by fastforce from assms(1) have "oalist_inv_raw ox xs'" by (simp add: xs oalist_inv_def) from assms(2) have "(k, v) \ set xs'" by (simp add: xs) show ?thesis by (simp add: xs, rule lookup_pair_eq_valueI, fact+) qed lemma lookup_raw_inj: assumes "oalist_inv (xs, ko)" and "oalist_inv (ys, ko)" and "lookup_raw (xs, ko) = lookup_raw (ys, ko)" shows "xs = ys" using assms unfolding lookup_raw.simps oalist_inv_alt by (rule lookup_pair_inj) subsubsection \@{const sorted_domain_raw}\ lemma set_sorted_domain_raw: assumes "oalist_inv xs" shows "set (sorted_domain_raw ko xs) = fst ` set (fst xs)" using assms by (simp add: sorted_domain_raw_def set_sort_oalist_aux) corollary in_sorted_domain_raw_iff_lookup_raw: assumes "oalist_inv xs" shows "k \ set (sorted_domain_raw ko xs) \ (lookup_raw xs k \ 0)" unfolding set_sorted_domain_raw[OF assms] proof - obtain xs' ko' where xs: "xs = (xs', ko')" by fastforce from assms show "k \ fst ` set (fst xs) \ (lookup_raw xs k \ 0)" by (simp add: xs oalist_inv_alt lookup_pair_eq_0) qed lemma sorted_sorted_domain_raw: assumes "oalist_inv xs" shows "sorted_wrt (lt_of_key_order (rep_key_order ko)) (sorted_domain_raw ko xs)" unfolding sorted_domain_raw_def oalist_inv_alt lt_of_key_order.rep_eq by (rule oalist_inv_rawD2, rule oalist_inv_raw_sort_oalist_aux, fact) subsubsection \@{const tl_raw}\ lemma oalist_inv_tl_raw: assumes "oalist_inv xs" shows "oalist_inv (tl_raw xs)" proof - obtain xs' ko where xs: "xs = (xs', ko)" by fastforce from assms show ?thesis unfolding xs tl_raw.simps oalist_inv_alt by (rule oalist_inv_raw_tl) qed lemma lookup_raw_tl_raw: assumes "oalist_inv xs" shows "lookup_raw (tl_raw xs) k = (if (\k'\fst ` set (fst xs). le (snd xs) k k') then 0 else lookup_raw xs k)" proof - obtain xs' ko where xs: "xs = (xs', ko)" by fastforce from assms show ?thesis by (simp add: xs lookup_pair_tl oalist_inv_alt split del: if_split cong: if_cong) qed lemma lookup_raw_tl_raw': assumes "oalist_inv xs" shows "lookup_raw (tl_raw xs) k = (if k = fst (List.hd (fst xs)) then 0 else lookup_raw xs k)" proof - obtain xs' ko where xs: "xs = (xs', ko)" by fastforce from assms show ?thesis by (simp add: xs lookup_pair_tl' oalist_inv_alt) qed subsubsection \@{const min_key_val_raw}\ lemma min_key_val_raw_alt: assumes "oalist_inv xs" and "fst xs \ []" shows "min_key_val_raw ko xs = List.hd (sort_oalist_aux ko xs)" proof - obtain xs' ox where xs: "xs = (xs', ox)" by fastforce from assms(2) have "xs' \ []" by (simp add: xs) interpret ko2: comparator2 "key_compare (rep_key_order ko)" "key_compare (rep_key_order ox)" .. from assms(1) have "oalist_inv_raw ox xs'" by (simp only: xs oalist_inv_alt) hence set_sort: "set (sort_oalist ko xs') = set xs'" by (rule ko2.set_sort_oalist) also from \xs' \ []\ have "... \ {}" by simp finally have "sort_oalist ko xs' \ []" by simp then obtain k v xs'' where eq: "sort_oalist ko xs' = (k, v) # xs''" using prod.exhaust list.exhaust by metis hence "(k, v) \ set xs'" by (simp add: set_sort[symmetric]) have *: "le ko k k'" if "k' \ fst ` set xs'" for k' proof - from that have "k' = k \ k' \ fst ` set xs''" by (simp add: set_sort[symmetric] eq) thus ?thesis proof assume "k' = k" thus ?thesis by simp next have "oalist_inv_raw ko ((k, v) # xs'')" unfolding eq[symmetric] by (fact oalist_inv_raw_sort_oalist) moreover assume "k' \ fst ` set xs''" ultimately have "lt ko k k'" by (rule oalist_inv_raw_ConsD3) thus ?thesis by simp qed qed from \xs' \ []\ have "min_list_param (\x y. le ko (fst x) (fst y)) xs' \ set xs'" by (rule min_list_param_in) with \oalist_inv_raw ox xs'\ have "lookup_pair ox xs' (fst (min_list_param (\x y. le ko (fst x) (fst y)) xs')) = snd (min_list_param (\x y. le ko (fst x) (fst y)) xs')" by (auto intro: lookup_pair_eq_valueI) moreover have 1: "fst (min_list_param (\x y. le ko (fst x) (fst y)) xs') = k" proof (rule antisym) from order_trans have "transp (\x y. le ko (fst x) (fst y))" by (rule transpI) hence "le ko (fst (min_list_param (\x y. le ko (fst x) (fst y)) xs')) (fst (k, v))" using linear \(k, v) \ set xs'\ by (rule min_list_param_minimal) thus "le ko (fst (min_list_param (\x y. le ko (fst x) (fst y)) xs')) k" by simp next show "le ko k (fst (min_list_param (\x y. le ko (fst x) (fst y)) xs'))" by (rule *, rule imageI, fact) qed ultimately have "snd (min_list_param (\x y. le ko (fst x) (fst y)) xs') = lookup_pair ox xs' k" by simp also from \oalist_inv_raw ox xs'\ \(k, v) \ set xs'\ have "... = v" by (rule lookup_pair_eq_valueI) finally have "snd (min_list_param (\x y. le ko (fst x) (fst y)) xs') = v" . with 1 have "min_list_param (\x y. le ko (fst x) (fst y)) xs' = (k, v)" by auto thus ?thesis by (simp add: xs eq) qed lemma min_key_val_raw_in: assumes "fst xs \ []" shows "min_key_val_raw ko xs \ set (fst xs)" proof - obtain xs' ox where xs: "xs = (xs', ox)" by fastforce from assms have "xs' \ []" by (simp add: xs) show ?thesis unfolding xs proof (simp, intro conjI impI) from \xs' \ []\ show "hd xs' \ set xs'" by simp next from \xs' \ []\ show "min_list_param (\x y. le ko (fst x) (fst y)) xs' \ set xs'" by (rule min_list_param_in) qed qed lemma snd_min_key_val_raw: assumes "oalist_inv xs" and "fst xs \ []" shows "snd (min_key_val_raw ko xs) = lookup_raw xs (fst (min_key_val_raw ko xs))" proof - obtain xs' ox where xs: "xs = (xs', ox)" by fastforce from assms(1) have "oalist_inv_raw ox xs'" by (simp only: xs oalist_inv_alt) from assms(2) have "min_key_val_raw ko xs \ set (fst xs)" by (rule min_key_val_raw_in) hence *: "min_key_val_raw ko (xs', ox) \ set xs'" by (simp add: xs) show ?thesis unfolding xs lookup_raw.simps by (rule HOL.sym, rule lookup_pair_eq_valueI, fact, simp add: * del: min_key_val_raw.simps) qed lemma min_key_val_raw_minimal: assumes "oalist_inv xs" and "z \ set (fst xs)" shows "le ko (fst (min_key_val_raw ko xs)) (fst z)" proof - obtain xs' ox where xs: "xs = (xs', ox)" by fastforce from assms have "oalist_inv (xs', ox)" and "z \ set xs'" by (simp_all add: xs) show ?thesis unfolding xs proof (simp, intro conjI impI) from \z \ set xs'\ have "xs' \ []" by auto then obtain k v ys where xs': "xs' = (k, v) # ys" using prod.exhaust list.exhaust by metis from \z \ set xs'\ have "z = (k, v) \ z \ set ys" by (simp add: xs') thus "le ox (fst (hd xs')) (fst z)" proof assume "z = (k, v)" show ?thesis by (simp add: xs' \z = (k, v)\) next assume "z \ set ys" hence "fst z \ fst ` set ys" by fastforce with \oalist_inv (xs', ox)\ have "lt ox k (fst z)" unfolding xs' oalist_inv_alt lt_of_key_order.rep_eq by (rule oalist_inv_raw_ConsD3) thus ?thesis by (simp add: xs') qed next show "le ko (fst (min_list_param (\x y. le ko (fst x) (fst y)) xs')) (fst z)" proof (rule min_list_param_minimal[of "\x y. le ko (fst x) (fst y)"]) thm trans local.trans order.trans local.order_trans print_context show "transp (\x y. le ko (fst x) (fst y))" by (metis (no_types, lifting) order_trans transpI) qed (auto intro: \z \ set xs'\) qed qed subsubsection \@{const filter_raw}\ lemma oalist_inv_filter_raw: assumes "oalist_inv xs" shows "oalist_inv (filter_raw P xs)" proof - obtain xs' ko where xs: "xs = (xs', ko)" by fastforce from assms show ?thesis unfolding xs filter_raw.simps oalist_inv_alt by (rule oalist_inv_raw_filter) qed lemma lookup_raw_filter_raw: assumes "oalist_inv xs" shows "lookup_raw (filter_raw P xs) k = (let v = lookup_raw xs k in if P (k, v) then v else 0)" proof - obtain xs' ko where xs: "xs = (xs', ko)" by fastforce from assms show ?thesis unfolding xs lookup_raw.simps filter_raw.simps oalist_inv_alt by (rule lookup_pair_filter) qed subsubsection \@{const update_by_raw}\ lemma oalist_inv_update_by_raw: assumes "oalist_inv xs" shows "oalist_inv (update_by_raw kv xs)" proof - obtain xs' ko where xs: "xs = (xs', ko)" by fastforce from assms show ?thesis unfolding xs update_by_raw.simps oalist_inv_alt by (rule oalist_inv_raw_update_by_pair) qed lemma lookup_raw_update_by_raw: assumes "oalist_inv xs" shows "lookup_raw (update_by_raw (k1, v) xs) k2 = (if k1 = k2 then v else lookup_raw xs k2)" proof - obtain xs' ko where xs: "xs = (xs', ko)" by fastforce from assms show ?thesis unfolding xs lookup_raw.simps update_by_raw.simps oalist_inv_alt by (rule lookup_pair_update_by_pair) qed subsubsection \@{const update_by_fun_raw} and @{const update_by_fun_gr_raw}\ lemma update_by_fun_raw_eq_update_by_raw: assumes "oalist_inv xs" shows "update_by_fun_raw k f xs = update_by_raw (k, f (lookup_raw xs k)) xs" proof - obtain xs' ko where xs: "xs = (xs', ko)" by fastforce from assms have "oalist_inv_raw ko xs'" by (simp only: xs oalist_inv_alt) show ?thesis unfolding xs update_by_fun_raw.simps lookup_raw.simps update_by_raw.simps by (rule, rule conjI, rule update_by_fun_pair_eq_update_by_pair, fact, fact HOL.refl) qed corollary oalist_inv_update_by_fun_raw: assumes "oalist_inv xs" shows "oalist_inv (update_by_fun_raw k f xs)" unfolding update_by_fun_raw_eq_update_by_raw[OF assms] using assms by (rule oalist_inv_update_by_raw) corollary lookup_raw_update_by_fun_raw: assumes "oalist_inv xs" shows "lookup_raw (update_by_fun_raw k1 f xs) k2 = (if k1 = k2 then f else id) (lookup_raw xs k2)" using assms by (simp add: update_by_fun_raw_eq_update_by_raw lookup_raw_update_by_raw) lemma update_by_fun_gr_raw_eq_update_by_fun_raw: assumes "oalist_inv xs" shows "update_by_fun_gr_raw k f xs = update_by_fun_raw k f xs" proof - obtain xs' ko where xs: "xs = (xs', ko)" by fastforce from assms have "oalist_inv_raw ko xs'" by (simp only: xs oalist_inv_alt) show ?thesis unfolding xs update_by_fun_raw.simps lookup_raw.simps update_by_fun_gr_raw.simps by (rule, rule conjI, rule update_by_fun_gr_pair_eq_update_by_fun_pair, fact, fact HOL.refl) qed corollary oalist_inv_update_by_fun_gr_raw: assumes "oalist_inv xs" shows "oalist_inv (update_by_fun_gr_raw k f xs)" unfolding update_by_fun_gr_raw_eq_update_by_fun_raw[OF assms] using assms by (rule oalist_inv_update_by_fun_raw) corollary lookup_raw_update_by_fun_gr_raw: assumes "oalist_inv xs" shows "lookup_raw (update_by_fun_gr_raw k1 f xs) k2 = (if k1 = k2 then f else id) (lookup_raw xs k2)" using assms by (simp add: update_by_fun_gr_raw_eq_update_by_fun_raw lookup_raw_update_by_fun_raw) subsubsection \@{const map_raw} and @{const map_val_raw}\ lemma map_raw_cong: assumes "\kv. kv \ set (fst xs) \ f kv = g kv" shows "map_raw f xs = map_raw g xs" proof - obtain xs' ko where xs: "xs = (xs', ko)" by fastforce from assms have "f kv = g kv" if "kv \ set xs'" for kv using that by (simp add: xs) thus ?thesis by (simp add: xs, rule map_pair_cong) qed lemma map_raw_subset: "set (fst (map_raw f xs)) \ f ` set (fst xs)" proof - obtain xs' ko where xs: "xs = (xs', ko)" by fastforce show ?thesis by (simp add: xs map_pair_subset) qed lemma oalist_inv_map_raw: assumes "oalist_inv xs" and "\a b. key_compare (rep_key_order (snd xs)) (fst (f a)) (fst (f b)) = key_compare (rep_key_order (snd xs)) (fst a) (fst b)" shows "oalist_inv (map_raw f xs)" proof - obtain xs' ko where xs: "xs = (xs', ko)" by fastforce from assms(1) have "oalist_inv (xs', ko)" by (simp only: xs) moreover from assms(2) have "\a b. key_compare (rep_key_order ko) (fst (f a)) (fst (f b)) = key_compare (rep_key_order ko) (fst a) (fst b)" by (simp add: xs) ultimately show ?thesis unfolding xs map_raw.simps oalist_inv_alt by (rule oalist_inv_raw_map_pair) qed lemma lookup_raw_map_raw: assumes "oalist_inv xs" and "snd (f (k, 0)) = 0" and "\a b. key_compare (rep_key_order (snd xs)) (fst (f a)) (fst (f b)) = key_compare (rep_key_order (snd xs)) (fst a) (fst b)" shows "lookup_raw (map_raw f xs) (fst (f (k, v))) = snd (f (k, lookup_raw xs k))" proof - obtain xs' ko where xs: "xs = (xs', ko)" by fastforce from assms(1) have "oalist_inv (xs', ko)" by (simp only: xs) moreover note assms(2) moreover from assms(3) have "\a b. key_compare (rep_key_order ko) (fst (f a)) (fst (f b)) = key_compare (rep_key_order ko) (fst a) (fst b)" by (simp add: xs) ultimately show ?thesis unfolding xs lookup_raw.simps map_raw.simps oalist_inv_alt by (rule lookup_pair_map_pair) qed lemma map_raw_id: assumes "oalist_inv xs" shows "map_raw id xs = xs" proof - obtain xs' ko where xs: "xs = (xs', ko)" by fastforce from assms have "oalist_inv_raw ko xs'" by (simp only: xs oalist_inv_alt) hence "map_pair id xs' = xs'" proof (induct xs' rule: oalist_inv_raw_induct) case Nil show ?case by simp next case (Cons k v xs') show ?case by (simp add: Let_def Cons(3, 5) id_def[symmetric]) qed thus ?thesis by (simp add: xs) qed lemma map_val_raw_cong: assumes "\k v. (k, v) \ set (fst xs) \ f k v = g k v" shows "map_val_raw f xs = map_val_raw g xs" proof (rule map_raw_cong) fix kv assume "kv \ set (fst xs)" moreover obtain k v where "kv = (k, v)" by fastforce ultimately show "(case kv of (k, v) \ (k, f k v)) = (case kv of (k, v) \ (k, g k v))" by (simp add: assms) qed lemma oalist_inv_map_val_raw: assumes "oalist_inv xs" shows "oalist_inv (map_val_raw f xs)" proof - obtain xs' ko where xs: "xs = (xs', ko)" by fastforce from assms show ?thesis unfolding xs map_raw.simps oalist_inv_alt by (rule oalist_inv_raw_map_val_pair) qed lemma lookup_raw_map_val_raw: assumes "oalist_inv xs" and "f k 0 = 0" shows "lookup_raw (map_val_raw f xs) k = f k (lookup_raw xs k)" proof - obtain xs' ko where xs: "xs = (xs', ko)" by fastforce from assms show ?thesis unfolding xs map_raw.simps lookup_raw.simps oalist_inv_alt by (rule lookup_pair_map_val_pair) qed subsubsection \@{const map2_val_raw}\ definition map2_val_compat' :: "(('a, 'b::zero, 'o) oalist_raw \ ('a, 'c::zero, 'o) oalist_raw) \ bool" where "map2_val_compat' f \ (\zs. (oalist_inv zs \ (oalist_inv (f zs) \ snd (f zs) = snd zs \ fst ` set (fst (f zs)) \ fst ` set (fst zs))))" lemma map2_val_compat'I: assumes "\zs. oalist_inv zs \ oalist_inv (f zs)" and "\zs. oalist_inv zs \ snd (f zs) = snd zs" and "\zs. oalist_inv zs \ fst ` set (fst (f zs)) \ fst ` set (fst zs)" shows "map2_val_compat' f" unfolding map2_val_compat'_def using assms by blast lemma map2_val_compat'D1: assumes "map2_val_compat' f" and "oalist_inv zs" shows "oalist_inv (f zs)" using assms unfolding map2_val_compat'_def by blast lemma map2_val_compat'D2: assumes "map2_val_compat' f" and "oalist_inv zs" shows "snd (f zs) = snd zs" using assms unfolding map2_val_compat'_def by blast lemma map2_val_compat'D3: assumes "map2_val_compat' f" and "oalist_inv zs" shows "fst ` set (fst (f zs)) \ fst ` set (fst zs)" using assms unfolding map2_val_compat'_def by blast lemma map2_val_compat'_map_val_raw: "map2_val_compat' (map_val_raw f)" proof (rule map2_val_compat'I, erule oalist_inv_map_val_raw) fix zs::"('a, 'b, 'o) oalist_raw" obtain zs' ko where "zs = (zs', ko)" by fastforce thus "snd (map_val_raw f zs) = snd zs" by simp next fix zs::"('a, 'b, 'o) oalist_raw" obtain zs' ko where zs: "zs = (zs', ko)" by fastforce show "fst ` set (fst (map_val_raw f zs)) \ fst ` set (fst zs)" proof (simp add: zs) from map_pair_subset have "fst ` set (map_val_pair f zs') \ fst ` (\(k, v). (k, f k v)) ` set zs'" by (rule image_mono) also have "... = fst ` set zs'" by force finally show "fst ` set (map_val_pair f zs') \ fst ` set zs'" . qed qed lemma map2_val_compat'_id: "map2_val_compat' id" by (rule map2_val_compat'I, auto) lemma map2_val_compat'_imp_map2_val_compat: assumes "map2_val_compat' g" shows "map2_val_compat ko (\zs. fst (g (zs, ko)))" proof (rule map2_val_compatI) fix zs::"('a \ 'b) list" assume a: "oalist_inv_raw ko zs" hence "oalist_inv (zs, ko)" by (simp only: oalist_inv_alt) with assms have "oalist_inv (g (zs, ko))" by (rule map2_val_compat'D1) hence "oalist_inv (fst (g (zs, ko)), snd (g (zs, ko)))" by simp thus "oalist_inv_raw ko (fst (g (zs, ko)))" using assms a by (simp add: oalist_inv_alt map2_val_compat'D2) next fix zs::"('a \ 'b) list" assume a: "oalist_inv_raw ko zs" hence "oalist_inv (zs, ko)" by (simp only: oalist_inv_alt) with assms have "fst ` set (fst (g (zs, ko))) \ fst ` set (fst (zs, ko))" by (rule map2_val_compat'D3) thus "fst ` set (fst (g (zs, ko))) \ fst ` set zs" by simp qed lemma oalist_inv_map2_val_raw: assumes "oalist_inv xs" and "oalist_inv ys" assumes "map2_val_compat' g" and "map2_val_compat' h" shows "oalist_inv (map2_val_raw f g h xs ys)" proof - obtain xs' ox where xs: "xs = (xs', ox)" by fastforce let ?ys = "sort_oalist_aux ox ys" from assms(1) have "oalist_inv_raw ox xs'" by (simp add: xs oalist_inv_alt) moreover from assms(2) have "oalist_inv_raw ox (sort_oalist_aux ox ys)" by (rule oalist_inv_raw_sort_oalist_aux) moreover from assms(3) have "map2_val_compat ko (\zs. fst (g (zs, ko)))" for ko by (rule map2_val_compat'_imp_map2_val_compat) moreover from assms(4) have "map2_val_compat ko (\zs. fst (h (zs, ko)))" for ko by (rule map2_val_compat'_imp_map2_val_compat) ultimately have "oalist_inv_raw ox (map2_val_pair ox f (\zs. fst (g (zs, ox))) (\zs. fst (h (zs, ox))) xs' ?ys)" by (rule oalist_inv_raw_map2_val_pair) thus ?thesis by (simp add: xs oalist_inv_alt) qed lemma lookup_raw_map2_val_raw: assumes "oalist_inv xs" and "oalist_inv ys" assumes "map2_val_compat' g" and "map2_val_compat' h" assumes "\zs. oalist_inv zs \ g zs = map_val_raw (\k v. f k v 0) zs" and "\zs. oalist_inv zs \ h zs = map_val_raw (\k. f k 0) zs" and "\k. f k 0 0 = 0" shows "lookup_raw (map2_val_raw f g h xs ys) k0 = f k0 (lookup_raw xs k0) (lookup_raw ys k0)" proof - obtain xs' ox where xs: "xs = (xs', ox)" by fastforce let ?ys = "sort_oalist_aux ox ys" from assms(1) have "oalist_inv_raw ox xs'" by (simp add: xs oalist_inv_alt) moreover from assms(2) have "oalist_inv_raw ox (sort_oalist_aux ox ys)" by (rule oalist_inv_raw_sort_oalist_aux) moreover from assms(3) have "map2_val_compat ko (\zs. fst (g (zs, ko)))" for ko by (rule map2_val_compat'_imp_map2_val_compat) moreover from assms(4) have "map2_val_compat ko (\zs. fst (h (zs, ko)))" for ko by (rule map2_val_compat'_imp_map2_val_compat) ultimately have "lookup_pair ox (map2_val_pair ox f (\zs. fst (g (zs, ox))) (\zs. fst (h (zs, ox))) xs' ?ys) k0 = f k0 (lookup_pair ox xs' k0) (lookup_pair ox ?ys k0)" using _ _ assms(7) proof (rule lookup_pair_map2_val_pair) fix zs::"('a \ 'b) list" assume "oalist_inv_raw ox zs" hence "oalist_inv (zs, ox)" by (simp only: oalist_inv_alt) hence "g (zs, ox) = map_val_raw (\k v. f k v 0) (zs, ox)" by (rule assms(5)) thus "fst (g (zs, ox)) = map_val_pair (\k v. f k v 0) zs" by simp next fix zs::"('a \ 'c) list" assume "oalist_inv_raw ox zs" hence "oalist_inv (zs, ox)" by (simp only: oalist_inv_alt) hence "h (zs, ox) = map_val_raw (\k. f k 0) (zs, ox)" by (rule assms(6)) thus "fst (h (zs, ox)) = map_val_pair (\k. f k 0) zs" by simp qed also from assms(2) have "... = f k0 (lookup_pair ox xs' k0) (lookup_raw ys k0)" by (simp only: lookup_pair_sort_oalist_aux) finally have *: "lookup_pair ox (map2_val_pair ox f (\zs. fst (g (zs, ox))) (\zs. fst (h (zs, ox))) xs' ?ys) k0 = f k0 (lookup_pair ox xs' k0) (lookup_raw ys k0)" . thus ?thesis by (simp add: xs) qed lemma map2_val_raw_singleton_eq_update_by_fun_raw: assumes "oalist_inv xs" assumes "\k x. f k x 0 = x" and "\zs. oalist_inv zs \ g zs = zs" and "\ko. h ([(k, v)], ko) = map_val_raw (\k. f k 0) ([(k, v)], ko)" shows "map2_val_raw f g h xs ([(k, v)], ko) = update_by_fun_raw k (\x. f k x v) xs" proof - obtain xs' ox where xs: "xs = (xs', ox)" by fastforce let ?ys = "sort_oalist ox [(k, v)]" from assms(1) have inv: "oalist_inv (xs', ox)" by (simp only: xs) hence inv_raw: "oalist_inv_raw ox xs'" by (simp only: oalist_inv_alt) show ?thesis proof (simp add: xs, intro conjI impI) assume "ox = ko" from inv_raw have "oalist_inv_raw ko xs'" by (simp only: \ox = ko\) thus "map2_val_pair ko f (\zs. fst (g (zs, ko))) (\zs. fst (h (zs, ko))) xs' [(k, v)] = update_by_fun_pair ko k (\x. f k x v) xs'" using assms(2) proof (rule map2_val_pair_singleton_eq_update_by_fun_pair) fix zs::"('a \ 'b) list" assume "oalist_inv_raw ko zs" hence "oalist_inv (zs, ko)" by (simp only: oalist_inv_alt) hence "g (zs, ko) = (zs, ko)" by (rule assms(3)) thus "fst (g (zs, ko)) = zs" by simp next show "fst (h ([(k, v)], ko)) = map_val_pair (\k. f k 0) [(k, v)]" by (simp add: assms(4)) qed next show "map2_val_pair ox f (\zs. fst (g (zs, ox))) (\zs. fst (h (zs, ox))) xs' (sort_oalist ox [(k, v)]) = update_by_fun_pair ox k (\x. f k x v) xs'" proof (cases "v = 0") case True have eq1: "sort_oalist ox [(k, 0)] = []" by (simp add: sort_oalist_def) from inv have eq2: "g (xs', ox) = (xs', ox)" by (rule assms(3)) show ?thesis by (simp add: True eq1 eq2 assms(2) update_by_fun_pair_eq_update_by_pair[OF inv_raw], rule HOL.sym, rule update_by_pair_id, fact inv_raw, fact HOL.refl) next case False hence "oalist_inv_raw ox [(k, v)]" by (simp add: oalist_inv_raw_singleton) hence eq: "sort_oalist ox [(k, v)] = [(k, v)]" by (rule sort_oalist_id) show ?thesis unfolding eq using inv_raw assms(2) proof (rule map2_val_pair_singleton_eq_update_by_fun_pair) fix zs::"('a \ 'b) list" assume "oalist_inv_raw ox zs" hence "oalist_inv (zs, ox)" by (simp only: oalist_inv_alt) hence "g (zs, ox) = (zs, ox)" by (rule assms(3)) thus "fst (g (zs, ox)) = zs" by simp next show "fst (h ([(k, v)], ox)) = map_val_pair (\k. f k 0) [(k, v)]" by (simp add: assms(4)) qed qed qed qed subsubsection \@{const lex_ord_raw}\ lemma lex_ord_raw_EqI: assumes "oalist_inv xs" and "oalist_inv ys" and "\k. k \ fst ` set (fst xs) \ fst ` set (fst ys) \ f k (lookup_raw xs k) (lookup_raw ys k) = Some Eq" shows "lex_ord_raw ko f xs ys = Some Eq" unfolding lex_ord_raw_def by (rule lex_ord_pair_EqI, simp_all add: assms oalist_inv_raw_sort_oalist_aux lookup_pair_sort_oalist_aux set_sort_oalist_aux) lemma lex_ord_raw_valI: assumes "oalist_inv xs" and "oalist_inv ys" and "aux \ Some Eq" assumes "k \ fst ` set (fst xs) \ fst ` set (fst ys)" and "aux = f k (lookup_raw xs k) (lookup_raw ys k)" and "\k'. k' \ fst ` set (fst xs) \ fst ` set (fst ys) \ lt ko k' k \ f k' (lookup_raw xs k') (lookup_raw ys k') = Some Eq" shows "lex_ord_raw ko f xs ys = aux" unfolding lex_ord_raw_def using oalist_inv_sort_oalist_aux[OF assms(1)] oalist_inv_raw_sort_oalist_aux[OF assms(2)] assms(3) unfolding oalist_inv_alt proof (rule lex_ord_pair_valI) from assms(1, 2, 4) show "k \ fst ` set (sort_oalist_aux ko xs) \ fst ` set (sort_oalist_aux ko ys)" by (simp add: set_sort_oalist_aux) next from assms(1, 2, 5) show "aux = f k (lookup_pair ko (sort_oalist_aux ko xs) k) (lookup_pair ko (sort_oalist_aux ko ys) k)" by (simp add: lookup_pair_sort_oalist_aux) next fix k' assume "k' \ fst ` set (sort_oalist_aux ko xs) \ fst ` set (sort_oalist_aux ko ys)" with assms(1, 2) have "k' \ fst ` set (fst xs) \ fst ` set (fst ys)" by (simp add: set_sort_oalist_aux) moreover assume "lt ko k' k" ultimately have "f k' (lookup_raw xs k') (lookup_raw ys k') = Some Eq" by (rule assms(6)) with assms(1, 2) show "f k' (lookup_pair ko (sort_oalist_aux ko xs) k') (lookup_pair ko (sort_oalist_aux ko ys) k') = Some Eq" by (simp add: lookup_pair_sort_oalist_aux) qed lemma lex_ord_raw_EqD: assumes "oalist_inv xs" and "oalist_inv ys" and "lex_ord_raw ko f xs ys = Some Eq" and "k \ fst ` set (fst xs) \ fst ` set (fst ys)" shows "f k (lookup_raw xs k) (lookup_raw ys k) = Some Eq" proof - have "f k (lookup_pair ko (sort_oalist_aux ko xs) k) (lookup_pair ko (sort_oalist_aux ko ys) k) = Some Eq" by (rule lex_ord_pair_EqD[where f=f], simp_all add: oalist_inv_raw_sort_oalist_aux assms lex_ord_raw_def[symmetric] set_sort_oalist_aux del: Un_iff) with assms(1, 2) show ?thesis by (simp add: lookup_pair_sort_oalist_aux) qed lemma lex_ord_raw_valE: assumes "oalist_inv xs" and "oalist_inv ys" and "lex_ord_raw ko f xs ys = aux" and "aux \ Some Eq" obtains k where "k \ fst ` set (fst xs) \ fst ` set (fst ys)" and "aux = f k (lookup_raw xs k) (lookup_raw ys k)" and "\k'. k' \ fst ` set (fst xs) \ fst ` set (fst ys) \ lt ko k' k \ f k' (lookup_raw xs k') (lookup_raw ys k') = Some Eq" proof - let ?xs = "sort_oalist_aux ko xs" let ?ys = "sort_oalist_aux ko ys" from assms(3) have "lex_ord_pair ko f ?xs ?ys = aux" by (simp only: lex_ord_raw_def) with oalist_inv_sort_oalist_aux[OF assms(1)] oalist_inv_sort_oalist_aux[OF assms(2)] obtain k where a: "k \ fst ` set ?xs \ fst ` set ?ys" and b: "aux = f k (lookup_pair ko ?xs k) (lookup_pair ko ?ys k)" and c: "\k'. k' \ fst ` set ?xs \ fst ` set ?ys \ lt ko k' k \ f k' (lookup_pair ko ?xs k') (lookup_pair ko ?ys k') = Some Eq" using assms(4) unfolding oalist_inv_alt by (rule lex_ord_pair_valE, blast) from a have "k \ fst ` set (fst xs) \ fst ` set (fst ys)" by (simp add: set_sort_oalist_aux assms(1, 2)) moreover from b have "aux = f k (lookup_raw xs k) (lookup_raw ys k)" by (simp add: lookup_pair_sort_oalist_aux assms(1, 2)) moreover have "f k' (lookup_raw xs k') (lookup_raw ys k') = Some Eq" if k'_in: "k' \ fst ` set (fst xs) \ fst ` set (fst ys)" and k'_less: "lt ko k' k" for k' proof - have "f k' (lookup_raw xs k') (lookup_raw ys k') = f k' (lookup_pair ko ?xs k') (lookup_pair ko ?ys k')" by (simp add: lookup_pair_sort_oalist_aux assms(1, 2)) also have "... = Some Eq" proof (rule c) from k'_in show "k' \ fst ` set ?xs \ fst ` set ?ys" by (simp add: set_sort_oalist_aux assms(1, 2)) next from k'_less show "lt ko k' k" by (simp only: lt_of_key_order.rep_eq) qed finally show ?thesis . qed ultimately show ?thesis .. qed subsubsection \@{const prod_ord_raw}\ lemma prod_ord_rawI: assumes "oalist_inv xs" and "oalist_inv ys" and "\k. k \ fst ` set (fst xs) \ fst ` set (fst ys) \ P k (lookup_raw xs k) (lookup_raw ys k)" shows "prod_ord_raw P xs ys" proof - obtain xs' ox where xs: "xs = (xs', ox)" by fastforce from assms(1) have "oalist_inv_raw ox xs'" by (simp only: xs oalist_inv_alt) hence *: "prod_ord_pair ox P xs' (sort_oalist_aux ox ys)" using oalist_inv_raw_sort_oalist_aux proof (rule prod_ord_pairI) fix k assume "k \ fst ` set xs' \ fst ` set (sort_oalist_aux ox ys)" hence "k \ fst ` set (fst xs) \ fst ` set (fst ys)" by (simp add: xs set_sort_oalist_aux assms(2)) hence "P k (lookup_raw xs k) (lookup_raw ys k)" by (rule assms(3)) thus "P k (lookup_pair ox xs' k) (lookup_pair ox (sort_oalist_aux ox ys) k)" by (simp add: xs lookup_pair_sort_oalist_aux assms(2)) qed fact thus ?thesis by (simp add: xs) qed lemma prod_ord_rawD: assumes "oalist_inv xs" and "oalist_inv ys" and "prod_ord_raw P xs ys" and "k \ fst ` set (fst xs) \ fst ` set (fst ys)" shows "P k (lookup_raw xs k) (lookup_raw ys k)" proof - obtain xs' ox where xs: "xs = (xs', ox)" by fastforce from assms(1) have "oalist_inv_raw ox xs'" by (simp only: xs oalist_inv_alt) moreover note oalist_inv_raw_sort_oalist_aux[OF assms(2)] moreover from assms(3) have "prod_ord_pair ox P xs' (sort_oalist_aux ox ys)" by (simp add: xs) moreover from assms(4) have "k \ fst ` set xs' \ fst ` set (sort_oalist_aux ox ys)" by (simp add: xs set_sort_oalist_aux assms(2)) ultimately have *: "P k (lookup_pair ox xs' k) (lookup_pair ox (sort_oalist_aux ox ys) k)" by (rule prod_ord_pairD) thus ?thesis by (simp add: xs lookup_pair_sort_oalist_aux assms(2)) qed corollary prod_ord_raw_alt: assumes "oalist_inv xs" and "oalist_inv ys" shows "prod_ord_raw P xs ys \ (\k\fst ` set (fst xs) \ fst ` set (fst ys). P k (lookup_raw xs k) (lookup_raw ys k))" using prod_ord_rawI[OF assms] prod_ord_rawD[OF assms] by meson subsubsection \@{const oalist_eq_raw}\ lemma oalist_eq_rawI: assumes "oalist_inv xs" and "oalist_inv ys" and "\k. k \ fst ` set (fst xs) \ fst ` set (fst ys) \ lookup_raw xs k = lookup_raw ys k" shows "oalist_eq_raw xs ys" proof - obtain xs' ox where xs: "xs = (xs', ox)" by fastforce from assms(1) have "oalist_inv_raw ox xs'" by (simp only: xs oalist_inv_alt) hence *: "xs' = sort_oalist_aux ox ys" using oalist_inv_raw_sort_oalist_aux[OF assms(2)] proof (rule lookup_pair_inj) show "lookup_pair ox xs' = lookup_pair ox (sort_oalist_aux ox ys)" proof fix k show "lookup_pair ox xs' k = lookup_pair ox (sort_oalist_aux ox ys) k" proof (cases "k \ fst ` set xs' \ fst ` set (sort_oalist_aux ox ys)") case True hence "k \ fst ` set (fst xs) \ fst ` set (fst ys)" by (simp add: xs set_sort_oalist_aux assms(2)) hence "lookup_raw xs k = lookup_raw ys k" by (rule assms(3)) thus ?thesis by (simp add: xs lookup_pair_sort_oalist_aux assms(2)) next case False hence "k \ fst ` set xs'" and "k \ fst ` set (sort_oalist_aux ox ys)" by simp_all with \oalist_inv_raw ox xs'\ oalist_inv_raw_sort_oalist_aux[OF assms(2)] have "lookup_pair ox xs' k = 0" and "lookup_pair ox (sort_oalist_aux ox ys) k = 0" by (simp_all add: lookup_pair_eq_0) thus ?thesis by simp qed qed qed thus ?thesis by (simp add: xs) qed lemma oalist_eq_rawD: assumes "oalist_inv ys" and "oalist_eq_raw xs ys" shows "lookup_raw xs = lookup_raw ys" proof - obtain xs' ox where xs: "xs = (xs', ox)" by fastforce from assms(2) have "xs' = sort_oalist_aux ox ys" by (simp add: xs) hence "lookup_pair ox xs' = lookup_pair ox (sort_oalist_aux ox ys)" by simp thus ?thesis by (simp add: xs lookup_pair_sort_oalist_aux assms(1)) qed lemma oalist_eq_raw_alt: assumes "oalist_inv xs" and "oalist_inv ys" shows "oalist_eq_raw xs ys \ (lookup_raw xs = lookup_raw ys)" using oalist_eq_rawI[OF assms] oalist_eq_rawD[OF assms(2)] by metis subsubsection \@{const sort_oalist_raw}\ lemma oalist_inv_sort_oalist_raw: "oalist_inv (sort_oalist_raw xs)" proof - obtain xs' ko where xs: "xs = (xs', ko)" by fastforce show ?thesis by (simp add: xs oalist_inv_raw_sort_oalist oalist_inv_alt) qed lemma sort_oalist_raw_id: assumes "oalist_inv xs" shows "sort_oalist_raw xs = xs" proof - obtain xs' ko where xs: "xs = (xs', ko)" by fastforce from assms have "oalist_inv_raw ko xs'" by (simp only: xs oalist_inv_alt) hence "sort_oalist ko xs' = xs'" by (rule sort_oalist_id) thus ?thesis by (simp add: xs) qed lemma set_sort_oalist_raw: assumes "distinct (map fst (fst xs))" shows "set (fst (sort_oalist_raw xs)) = {kv. kv \ set (fst xs) \ snd kv \ 0}" proof - obtain xs' ko where xs: "xs = (xs', ko)" by fastforce from assms have "distinct (map fst xs')" by (simp add: xs) hence "set (sort_oalist ko xs') = {kv \ set xs'. snd kv \ 0}" by (rule set_sort_oalist) thus ?thesis by (simp add: xs) qed end (* oalist_raw *) subsection \Fundamental Operations on One List\ locale oalist_abstract = oalist_raw rep_key_order for rep_key_order::"'o \ 'a key_order" + fixes list_of_oalist :: "'x \ ('a, 'b::zero, 'o) oalist_raw" fixes oalist_of_list :: "('a, 'b, 'o) oalist_raw \ 'x" assumes oalist_inv_list_of_oalist: "oalist_inv (list_of_oalist x)" and list_of_oalist_of_list: "list_of_oalist (oalist_of_list xs) = sort_oalist_raw xs" and oalist_of_list_of_oalist: "oalist_of_list (list_of_oalist x) = x" begin lemma list_of_oalist_of_list_id: assumes "oalist_inv xs" shows "list_of_oalist (oalist_of_list xs) = xs" proof - obtain xs' ox where xs: "xs = (xs', ox)" by fastforce from assms show ?thesis by (simp add: xs list_of_oalist_of_list sort_oalist_id oalist_inv_alt) qed definition lookup :: "'x \ 'a \ 'b" where "lookup xs = lookup_raw (list_of_oalist xs)" definition sorted_domain :: "'o \ 'x \ 'a list" where "sorted_domain ko xs = sorted_domain_raw ko (list_of_oalist xs)" definition empty :: "'o \ 'x" where "empty ko = oalist_of_list ([], ko)" definition reorder :: "'o \ 'x \ 'x" where "reorder ko xs = oalist_of_list (sort_oalist_aux ko (list_of_oalist xs), ko)" definition tl :: "'x \ 'x" where "tl xs = oalist_of_list (tl_raw (list_of_oalist xs))" definition hd :: "'x \ ('a \ 'b)" where "hd xs = List.hd (fst (list_of_oalist xs))" definition except_min :: "'o \ 'x \ 'x" where "except_min ko xs = tl (reorder ko xs)" definition min_key_val :: "'o \ 'x \ ('a \ 'b)" where "min_key_val ko xs = min_key_val_raw ko (list_of_oalist xs)" definition insert :: "('a \ 'b) \ 'x \ 'x" where "insert x xs = oalist_of_list (update_by_raw x (list_of_oalist xs))" definition update_by_fun :: "'a \ ('b \ 'b) \ 'x \ 'x" where "update_by_fun k f xs = oalist_of_list (update_by_fun_raw k f (list_of_oalist xs))" definition update_by_fun_gr :: "'a \ ('b \ 'b) \ 'x \ 'x" where "update_by_fun_gr k f xs = oalist_of_list (update_by_fun_gr_raw k f (list_of_oalist xs))" definition filter :: "(('a \ 'b) \ bool) \ 'x \ 'x" where "filter P xs = oalist_of_list (filter_raw P (list_of_oalist xs))" definition map2_val_neutr :: "('a \ 'b \ 'b \ 'b) \ 'x \ 'x \ 'x" where "map2_val_neutr f xs ys = oalist_of_list (map2_val_raw f id id (list_of_oalist xs) (list_of_oalist ys))" definition oalist_eq :: "'x \ 'x \ bool" where "oalist_eq xs ys = oalist_eq_raw (list_of_oalist xs) (list_of_oalist ys)" subsubsection \Invariant\ lemma zero_notin_list_of_oalist: "0 \ snd ` set (fst (list_of_oalist xs))" proof - from oalist_inv_list_of_oalist have "oalist_inv_raw (snd (list_of_oalist xs)) (fst (list_of_oalist xs))" by (simp only: oalist_inv_def) thus ?thesis by (rule oalist_inv_rawD1) qed lemma list_of_oalist_sorted: "sorted_wrt (lt (snd (list_of_oalist xs))) (map fst (fst (list_of_oalist xs)))" proof - from oalist_inv_list_of_oalist have "oalist_inv_raw (snd (list_of_oalist xs)) (fst (list_of_oalist xs))" by (simp only: oalist_inv_def) thus ?thesis by (rule oalist_inv_rawD2) qed subsubsection \@{const lookup}\ lemma lookup_eq_value: "v \ 0 \ lookup xs k = v \ ((k, v) \ set (fst (list_of_oalist xs)))" unfolding lookup_def using oalist_inv_list_of_oalist by (rule lookup_raw_eq_value) lemma lookup_eq_valueI: "(k, v) \ set (fst (list_of_oalist xs)) \ lookup xs k = v" unfolding lookup_def using oalist_inv_list_of_oalist by (rule lookup_raw_eq_valueI) lemma lookup_oalist_of_list: "distinct (map fst xs) \ lookup (oalist_of_list (xs, ko)) = lookup_dflt xs" by (simp add: list_of_oalist_of_list lookup_def lookup_pair_sort_oalist') subsubsection \@{const sorted_domain}\ lemma set_sorted_domain: "set (sorted_domain ko xs) = fst ` set (fst (list_of_oalist xs))" unfolding sorted_domain_def using oalist_inv_list_of_oalist by (rule set_sorted_domain_raw) lemma in_sorted_domain_iff_lookup: "k \ set (sorted_domain ko xs) \ (lookup xs k \ 0)" unfolding sorted_domain_def lookup_def using oalist_inv_list_of_oalist by (rule in_sorted_domain_raw_iff_lookup_raw) lemma sorted_sorted_domain: "sorted_wrt (lt ko) (sorted_domain ko xs)" unfolding sorted_domain_def lt_of_key_order.rep_eq[symmetric] using oalist_inv_list_of_oalist by (rule sorted_sorted_domain_raw) subsubsection \@{const empty} and Singletons\ lemma list_of_oalist_empty [simp, code abstract]: "list_of_oalist (empty ko) = ([], ko)" by (simp add: empty_def sort_oalist_def list_of_oalist_of_list) lemma lookup_empty: "lookup (empty ko) k = 0" by (simp add: lookup_def) lemma lookup_oalist_of_list_single: "lookup (oalist_of_list ([(k, v)], ko)) k' = (if k = k' then v else 0)" by (simp add: lookup_def list_of_oalist_of_list sort_oalist_def key_compare_Eq split: order.split) subsubsection \@{const reorder}\ lemma list_of_oalist_reorder [simp, code abstract]: "list_of_oalist (reorder ko xs) = (sort_oalist_aux ko (list_of_oalist xs), ko)" unfolding reorder_def by (rule list_of_oalist_of_list_id, simp add: oalist_inv_def, rule oalist_inv_raw_sort_oalist_aux, fact oalist_inv_list_of_oalist) lemma lookup_reorder: "lookup (reorder ko xs) k = lookup xs k" by (simp add: lookup_def lookup_pair_sort_oalist_aux oalist_inv_list_of_oalist) subsubsection \@{const hd} and @{const tl}\ lemma list_of_oalist_tl [simp, code abstract]: "list_of_oalist (tl xs) = tl_raw (list_of_oalist xs)" unfolding tl_def by (rule list_of_oalist_of_list_id, rule oalist_inv_tl_raw, fact oalist_inv_list_of_oalist) lemma lookup_tl: "lookup (tl xs) k = (if (\k'\fst ` set (fst (list_of_oalist xs)). le (snd (list_of_oalist xs)) k k') then 0 else lookup xs k)" by (simp add: lookup_def lookup_raw_tl_raw oalist_inv_list_of_oalist) lemma hd_in: assumes "fst (list_of_oalist xs) \ []" shows "hd xs \ set (fst (list_of_oalist xs))" unfolding hd_def using assms by (rule hd_in_set) lemma snd_hd: assumes "fst (list_of_oalist xs) \ []" shows "snd (hd xs) = lookup xs (fst (hd xs))" proof - from assms have *: "hd xs \ set (fst (list_of_oalist xs))" by (rule hd_in) show ?thesis by (rule HOL.sym, rule lookup_eq_valueI, simp add: *) qed lemma lookup_tl': "lookup (tl xs) k = (if k = fst (hd xs) then 0 else lookup xs k)" by (simp add: lookup_def lookup_raw_tl_raw' oalist_inv_list_of_oalist hd_def) lemma hd_tl: assumes "fst (list_of_oalist xs) \ []" shows "list_of_oalist xs = ((hd xs) # (fst (list_of_oalist (tl xs))), snd (list_of_oalist (tl xs)))" proof - obtain xs' ko where xs: "list_of_oalist xs = (xs', ko)" by fastforce from assms obtain x xs'' where xs': "xs' = x # xs''" unfolding xs fst_conv using list.exhaust by blast show ?thesis by (simp add: xs xs' hd_def) qed subsubsection \@{const min_key_val}\ lemma min_key_val_alt: assumes "fst (list_of_oalist xs) \ []" shows "min_key_val ko xs = hd (reorder ko xs)" using assms oalist_inv_list_of_oalist by (simp add: min_key_val_def hd_def min_key_val_raw_alt) lemma min_key_val_in: assumes "fst (list_of_oalist xs) \ []" shows "min_key_val ko xs \ set (fst (list_of_oalist xs))" unfolding min_key_val_def using assms by (rule min_key_val_raw_in) lemma snd_min_key_val: assumes "fst (list_of_oalist xs) \ []" shows "snd (min_key_val ko xs) = lookup xs (fst (min_key_val ko xs))" unfolding lookup_def min_key_val_def using oalist_inv_list_of_oalist assms by (rule snd_min_key_val_raw) lemma min_key_val_minimal: assumes "z \ set (fst (list_of_oalist xs))" shows "le ko (fst (min_key_val ko xs)) (fst z)" unfolding min_key_val_def by (rule min_key_val_raw_minimal, fact oalist_inv_list_of_oalist, fact) subsubsection \@{const except_min}\ lemma list_of_oalist_except_min [simp, code abstract]: "list_of_oalist (except_min ko xs) = (List.tl (sort_oalist_aux ko (list_of_oalist xs)), ko)" by (simp add: except_min_def) lemma except_min_Nil: assumes "fst (list_of_oalist xs) = []" shows "fst (list_of_oalist (except_min ko xs)) = []" proof - obtain xs' ox where eq: "list_of_oalist xs = (xs', ox)" by fastforce from assms have "xs' = []" by (simp add: eq) show ?thesis by (simp add: eq \xs' = []\ sort_oalist_def) qed lemma lookup_except_min: "lookup (except_min ko xs) k = (if (\k'\fst ` set (fst (list_of_oalist xs)). le ko k k') then 0 else lookup xs k)" by (simp add: except_min_def lookup_tl set_sort_oalist_aux oalist_inv_list_of_oalist lookup_reorder) lemma lookup_except_min': "lookup (except_min ko xs) k = (if k = fst (min_key_val ko xs) then 0 else lookup xs k)" proof (cases "fst (list_of_oalist xs) = []") case True hence "lookup xs k = 0" by (metis empty_def lookup_empty oalist_of_list_of_oalist prod.collapse) thus ?thesis by (simp add: lookup_except_min True) next case False thus ?thesis by (simp add: except_min_def lookup_tl' min_key_val_alt lookup_reorder) qed subsubsection \@{const insert}\ lemma list_of_oalist_insert [simp, code abstract]: "list_of_oalist (insert x xs) = update_by_raw x (list_of_oalist xs)" unfolding insert_def by (rule list_of_oalist_of_list_id, rule oalist_inv_update_by_raw, fact oalist_inv_list_of_oalist) lemma lookup_insert: "lookup (insert (k, v) xs) k' = (if k = k' then v else lookup xs k')" by (simp add: lookup_def lookup_raw_update_by_raw oalist_inv_list_of_oalist split del: if_split cong: if_cong) subsubsection \@{const update_by_fun} and @{const update_by_fun_gr}\ lemma list_of_oalist_update_by_fun [simp, code abstract]: "list_of_oalist (update_by_fun k f xs) = update_by_fun_raw k f (list_of_oalist xs)" unfolding update_by_fun_def by (rule list_of_oalist_of_list_id, rule oalist_inv_update_by_fun_raw, fact oalist_inv_list_of_oalist) lemma lookup_update_by_fun: "lookup (update_by_fun k f xs) k' = (if k = k' then f else id) (lookup xs k')" by (simp add: lookup_def lookup_raw_update_by_fun_raw oalist_inv_list_of_oalist split del: if_split cong: if_cong) lemma list_of_oalist_update_by_fun_gr [simp, code abstract]: "list_of_oalist (update_by_fun_gr k f xs) = update_by_fun_gr_raw k f (list_of_oalist xs)" unfolding update_by_fun_gr_def by (rule list_of_oalist_of_list_id, rule oalist_inv_update_by_fun_gr_raw, fact oalist_inv_list_of_oalist) lemma update_by_fun_gr_eq_update_by_fun: "update_by_fun_gr = update_by_fun" by (rule, rule, rule, simp add: update_by_fun_gr_def update_by_fun_def update_by_fun_gr_raw_eq_update_by_fun_raw oalist_inv_list_of_oalist) subsubsection \@{const filter}\ lemma list_of_oalist_filter [simp, code abstract]: "list_of_oalist (filter P xs) = filter_raw P (list_of_oalist xs)" unfolding filter_def by (rule list_of_oalist_of_list_id, rule oalist_inv_filter_raw, fact oalist_inv_list_of_oalist) lemma lookup_filter: "lookup (filter P xs) k = (let v = lookup xs k in if P (k, v) then v else 0)" by (simp add: lookup_def lookup_raw_filter_raw oalist_inv_list_of_oalist) subsubsection \@{const map2_val_neutr}\ lemma list_of_oalist_map2_val_neutr [simp, code abstract]: "list_of_oalist (map2_val_neutr f xs ys) = map2_val_raw f id id (list_of_oalist xs) (list_of_oalist ys)" unfolding map2_val_neutr_def by (rule list_of_oalist_of_list_id, rule oalist_inv_map2_val_raw, fact oalist_inv_list_of_oalist, fact oalist_inv_list_of_oalist, fact map2_val_compat'_id, fact map2_val_compat'_id) lemma lookup_map2_val_neutr: assumes "\k x. f k x 0 = x" and "\k x. f k 0 x = x" shows "lookup (map2_val_neutr f xs ys) k = f k (lookup xs k) (lookup ys k)" proof (simp add: lookup_def, rule lookup_raw_map2_val_raw) fix zs::"('a, 'b, 'o) oalist_raw" assume "oalist_inv zs" thus "id zs = map_val_raw (\k v. f k v 0) zs" by (simp add: assms(1) map_raw_id) next fix zs::"('a, 'b, 'o) oalist_raw" assume "oalist_inv zs" thus "id zs = map_val_raw (\k. f k 0) zs" by (simp add: assms(2) map_raw_id) qed (fact oalist_inv_list_of_oalist, fact oalist_inv_list_of_oalist, fact map2_val_compat'_id, fact map2_val_compat'_id, simp only: assms(1)) subsubsection \@{const oalist_eq}\ lemma oalist_eq_alt: "oalist_eq xs ys \ (lookup xs = lookup ys)" by (simp add: oalist_eq_def lookup_def oalist_eq_raw_alt oalist_inv_list_of_oalist) end (* oalist_abstract *) subsection \Fundamental Operations on Three Lists\ locale oalist_abstract3 = oalist_abstract rep_key_order list_of_oalistx oalist_of_listx + oay: oalist_abstract rep_key_order list_of_oalisty oalist_of_listy + oaz: oalist_abstract rep_key_order list_of_oalistz oalist_of_listz for rep_key_order :: "'o \ 'a key_order" and list_of_oalistx :: "'x \ ('a, 'b::zero, 'o) oalist_raw" and oalist_of_listx :: "('a, 'b, 'o) oalist_raw \ 'x" and list_of_oalisty :: "'y \ ('a, 'c::zero, 'o) oalist_raw" and oalist_of_listy :: "('a, 'c, 'o) oalist_raw \ 'y" and list_of_oalistz :: "'z \ ('a, 'd::zero, 'o) oalist_raw" and oalist_of_listz :: "('a, 'd, 'o) oalist_raw \ 'z" begin definition map_val :: "('a \ 'b \ 'c) \ 'x \ 'y" where "map_val f xs = oalist_of_listy (map_val_raw f (list_of_oalistx xs))" definition map2_val :: "('a \ 'b \ 'c \ 'd) \ 'x \ 'y \ 'z" where "map2_val f xs ys = oalist_of_listz (map2_val_raw f (map_val_raw (\k b. f k b 0)) (map_val_raw (\k. f k 0)) (list_of_oalistx xs) (list_of_oalisty ys))" definition map2_val_rneutr :: "('a \ 'b \ 'c \ 'b) \ 'x \ 'y \ 'x" where "map2_val_rneutr f xs ys = oalist_of_listx (map2_val_raw f id (map_val_raw (\k. f k 0)) (list_of_oalistx xs) (list_of_oalisty ys))" definition lex_ord :: "'o \ ('a \ ('b, 'c) comp_opt) \ ('x, 'y) comp_opt" where "lex_ord ko f xs ys = lex_ord_raw ko f (list_of_oalistx xs) (list_of_oalisty ys)" definition prod_ord :: "('a \ 'b \ 'c \ bool) \ 'x \ 'y \ bool" where "prod_ord f xs ys = prod_ord_raw f (list_of_oalistx xs) (list_of_oalisty ys)" subsubsection \@{const map_val}\ lemma map_val_cong: assumes "\k v. (k, v) \ set (fst (list_of_oalistx xs)) \ f k v = g k v" shows "map_val f xs = map_val g xs" unfolding map_val_def by (rule arg_cong[where f=oalist_of_listy], rule map_val_raw_cong, elim assms) lemma list_of_oalist_map_val [simp, code abstract]: "list_of_oalisty (map_val f xs) = map_val_raw f (list_of_oalistx xs)" unfolding map_val_def by (rule oay.list_of_oalist_of_list_id, rule oalist_inv_map_val_raw, fact oalist_inv_list_of_oalist) lemma lookup_map_val: "f k 0 = 0 \ oay.lookup (map_val f xs) k = f k (lookup xs k)" by (simp add: oay.lookup_def lookup_def lookup_raw_map_val_raw oalist_inv_list_of_oalist) subsubsection \@{const map2_val} and @{const map2_val_rneutr}\ lemma list_of_oalist_map2_val [simp, code abstract]: "list_of_oalistz (map2_val f xs ys) = map2_val_raw f (map_val_raw (\k b. f k b 0)) (map_val_raw (\k. f k 0)) (list_of_oalistx xs) (list_of_oalisty ys)" unfolding map2_val_def by (rule oaz.list_of_oalist_of_list_id, rule oalist_inv_map2_val_raw, fact oalist_inv_list_of_oalist, fact oay.oalist_inv_list_of_oalist, fact map2_val_compat'_map_val_raw, fact map2_val_compat'_map_val_raw) lemma list_of_oalist_map2_val_rneutr [simp, code abstract]: "list_of_oalistx (map2_val_rneutr f xs ys) = map2_val_raw f id (map_val_raw (\k c. f k 0 c)) (list_of_oalistx xs) (list_of_oalisty ys)" unfolding map2_val_rneutr_def by (rule list_of_oalist_of_list_id, rule oalist_inv_map2_val_raw, fact oalist_inv_list_of_oalist, fact oay.oalist_inv_list_of_oalist, fact map2_val_compat'_id, fact map2_val_compat'_map_val_raw) lemma lookup_map2_val: assumes "\k. f k 0 0 = 0" shows "oaz.lookup (map2_val f xs ys) k = f k (lookup xs k) (oay.lookup ys k)" by (simp add: oaz.lookup_def oay.lookup_def lookup_def lookup_raw_map2_val_raw map2_val_compat'_map_val_raw assms oalist_inv_list_of_oalist oay.oalist_inv_list_of_oalist) lemma lookup_map2_val_rneutr: assumes "\k x. f k x 0 = x" shows "lookup (map2_val_rneutr f xs ys) k = f k (lookup xs k) (oay.lookup ys k)" proof (simp add: lookup_def oay.lookup_def, rule lookup_raw_map2_val_raw) fix zs::"('a, 'b, 'o) oalist_raw" assume "oalist_inv zs" thus "id zs = map_val_raw (\k v. f k v 0) zs" by (simp add: assms map_raw_id) qed (fact oalist_inv_list_of_oalist, fact oay.oalist_inv_list_of_oalist, fact map2_val_compat'_id, fact map2_val_compat'_map_val_raw, rule HOL.refl, simp only: assms) lemma map2_val_rneutr_singleton_eq_update_by_fun: assumes "\a x. f a x 0 = x" and "list_of_oalisty ys = ([(k, v)], oy)" shows "map2_val_rneutr f xs ys = update_by_fun k (\x. f k x v) xs" by (simp add: map2_val_rneutr_def update_by_fun_def assms map2_val_raw_singleton_eq_update_by_fun_raw oalist_inv_list_of_oalist) subsubsection \@{const lex_ord} and @{const prod_ord}\ lemma lex_ord_EqI: "(\k. k \ fst ` set (fst (list_of_oalistx xs)) \ fst ` set (fst (list_of_oalisty ys)) \ f k (lookup xs k) (oay.lookup ys k) = Some Eq) \ lex_ord ko f xs ys = Some Eq" by (simp add: lex_ord_def lookup_def oay.lookup_def, rule lex_ord_raw_EqI, rule oalist_inv_list_of_oalist, rule oay.oalist_inv_list_of_oalist, blast) lemma lex_ord_valI: assumes "aux \ Some Eq" and "k \ fst ` set (fst (list_of_oalistx xs)) \ fst ` set (fst (list_of_oalisty ys))" shows "aux = f k (lookup xs k) (oay.lookup ys k) \ (\k'. k' \ fst ` set (fst (list_of_oalistx xs)) \ fst ` set (fst (list_of_oalisty ys)) \ lt ko k' k \ f k' (lookup xs k') (oay.lookup ys k') = Some Eq) \ lex_ord ko f xs ys = aux" by (simp (no_asm_use) add: lex_ord_def lookup_def oay.lookup_def, rule lex_ord_raw_valI, rule oalist_inv_list_of_oalist, rule oay.oalist_inv_list_of_oalist, rule assms(1), rule assms(2), blast+) lemma lex_ord_EqD: "lex_ord ko f xs ys = Some Eq \ k \ fst ` set (fst (list_of_oalistx xs)) \ fst ` set (fst (list_of_oalisty ys)) \ f k (lookup xs k) (oay.lookup ys k) = Some Eq" by (simp add: lex_ord_def lookup_def oay.lookup_def, rule lex_ord_raw_EqD[where f=f], rule oalist_inv_list_of_oalist, rule oay.oalist_inv_list_of_oalist, assumption, simp) lemma lex_ord_valE: assumes "lex_ord ko f xs ys = aux" and "aux \ Some Eq" obtains k where "k \ fst ` set (fst (list_of_oalistx xs)) \ fst ` set (fst (list_of_oalisty ys))" and "aux = f k (lookup xs k) (oay.lookup ys k)" and "\k'. k' \ fst ` set (fst (list_of_oalistx xs)) \ fst ` set (fst (list_of_oalisty ys)) \ lt ko k' k \ f k' (lookup xs k') (oay.lookup ys k') = Some Eq" proof - note oalist_inv_list_of_oalist oay.oalist_inv_list_of_oalist moreover from assms(1) have "lex_ord_raw ko f (list_of_oalistx xs) (list_of_oalisty ys) = aux" by (simp only: lex_ord_def) ultimately obtain k where 1: "k \ fst ` set (fst (list_of_oalistx xs)) \ fst ` set (fst (list_of_oalisty ys))" and "aux = f k (lookup_raw (list_of_oalistx xs) k) (lookup_raw (list_of_oalisty ys) k)" and "\k'. k' \ fst ` set (fst (list_of_oalistx xs)) \ fst ` set (fst (list_of_oalisty ys)) \ lt ko k' k \ f k' (lookup_raw (list_of_oalistx xs) k') (lookup_raw (list_of_oalisty ys) k') = Some Eq" using assms(2) by (rule lex_ord_raw_valE, blast) from this(2, 3) have "aux = f k (lookup xs k) (oay.lookup ys k)" and "\k'. k' \ fst ` set (fst (list_of_oalistx xs)) \ fst ` set (fst (list_of_oalisty ys)) \ lt ko k' k \ f k' (lookup xs k') (oay.lookup ys k') = Some Eq" by (simp_all only: lookup_def oay.lookup_def) with 1 show ?thesis .. qed lemma prod_ord_alt: "prod_ord P xs ys \ (\k\fst ` set (fst (list_of_oalistx xs)) \ fst ` set (fst (list_of_oalisty ys)). P k (lookup xs k) (oay.lookup ys k))" by (simp add: prod_ord_def lookup_def oay.lookup_def prod_ord_raw_alt oalist_inv_list_of_oalist oay.oalist_inv_list_of_oalist) end (* oalist_abstract3 *) subsection \Type \oalist\\ global_interpretation ko: comparator "key_compare ko" defines lookup_pair_ko = ko.lookup_pair and update_by_pair_ko = ko.update_by_pair and update_by_fun_pair_ko = ko.update_by_fun_pair and update_by_fun_gr_pair_ko = ko.update_by_fun_gr_pair and map2_val_pair_ko = ko.map2_val_pair and lex_ord_pair_ko = ko.lex_ord_pair and prod_ord_pair_ko = ko.prod_ord_pair and sort_oalist_ko' = ko.sort_oalist by (fact comparator_key_compare) lemma ko_le: "ko.le = le_of_key_order" by (intro ext, simp add: le_of_comp_def le_of_key_order_alt split: order.split) global_interpretation ko: oalist_raw "\x. x" rewrites "comparator.lookup_pair (key_compare ko) = lookup_pair_ko ko" and "comparator.update_by_pair (key_compare ko) = update_by_pair_ko ko" and "comparator.update_by_fun_pair (key_compare ko) = update_by_fun_pair_ko ko" and "comparator.update_by_fun_gr_pair (key_compare ko) = update_by_fun_gr_pair_ko ko" and "comparator.map2_val_pair (key_compare ko) = map2_val_pair_ko ko" and "comparator.lex_ord_pair (key_compare ko) = lex_ord_pair_ko ko" and "comparator.prod_ord_pair (key_compare ko) = prod_ord_pair_ko ko" and "comparator.sort_oalist (key_compare ko) = sort_oalist_ko' ko" defines sort_oalist_aux_ko = ko.sort_oalist_aux and lookup_ko = ko.lookup_raw and sorted_domain_ko = ko.sorted_domain_raw and tl_ko = ko.tl_raw and min_key_val_ko = ko.min_key_val_raw and update_by_ko = ko.update_by_raw and update_by_fun_ko = ko.update_by_fun_raw and update_by_fun_gr_ko = ko.update_by_fun_gr_raw and map2_val_ko = ko.map2_val_raw and lex_ord_ko = ko.lex_ord_raw and prod_ord_ko = ko.prod_ord_raw and oalist_eq_ko = ko.oalist_eq_raw and sort_oalist_ko = ko.sort_oalist_raw subgoal by (simp only: lookup_pair_ko_def) subgoal by (simp only: update_by_pair_ko_def) subgoal by (simp only: update_by_fun_pair_ko_def) subgoal by (simp only: update_by_fun_gr_pair_ko_def) subgoal by (simp only: map2_val_pair_ko_def) subgoal by (simp only: lex_ord_pair_ko_def) subgoal by (simp only: prod_ord_pair_ko_def) subgoal by (simp only: sort_oalist_ko'_def) done typedef (overloaded) ('a, 'b) oalist = "{xs::('a, 'b::zero, 'a key_order) oalist_raw. ko.oalist_inv xs}" morphisms list_of_oalist Abs_oalist by (auto simp: ko.oalist_inv_def intro: ko.oalist_inv_raw_Nil) lemma oalist_eq_iff: "xs = ys \ list_of_oalist xs = list_of_oalist ys" by (simp add: list_of_oalist_inject) lemma oalist_eqI: "list_of_oalist xs = list_of_oalist ys \ xs = ys" by (simp add: oalist_eq_iff) text \Formal, totalized constructor for @{typ "('a, 'b) oalist"}:\ definition OAlist :: "('a \ 'b) list \ 'a key_order \ ('a, 'b::zero) oalist" where "OAlist xs = Abs_oalist (sort_oalist_ko xs)" definition "oalist_of_list = OAlist" lemma oalist_inv_list_of_oalist: "ko.oalist_inv (list_of_oalist xs)" using list_of_oalist [of xs] by simp lemma list_of_oalist_OAlist: "list_of_oalist (OAlist xs) = sort_oalist_ko xs" proof - obtain xs' ox where xs: "xs = (xs', ox)" by fastforce show ?thesis by (simp add: xs OAlist_def Abs_oalist_inverse ko.oalist_inv_raw_sort_oalist ko.oalist_inv_alt) qed lemma OAlist_list_of_oalist [code abstype]: "OAlist (list_of_oalist xs) = xs" proof - obtain xs' ox where xs: "list_of_oalist xs = (xs', ox)" by fastforce have "ko.oalist_inv_raw ox xs'" by (simp add: xs[symmetric] ko.oalist_inv_alt[symmetric] oalist_inv_list_of_oalist) thus ?thesis by (simp add: xs OAlist_def ko.sort_oalist_id, simp add: list_of_oalist_inverse xs[symmetric]) qed lemma [code abstract]: "list_of_oalist (oalist_of_list xs) = sort_oalist_ko xs" by (simp add: list_of_oalist_OAlist oalist_of_list_def) global_interpretation oa: oalist_abstract "\x. x" list_of_oalist OAlist defines OAlist_lookup = oa.lookup and OAlist_sorted_domain = oa.sorted_domain and OAlist_empty = oa.empty and OAlist_reorder = oa.reorder and OAlist_tl = oa.tl and OAlist_hd = oa.hd and OAlist_except_min = oa.except_min and OAlist_min_key_val = oa.min_key_val and OAlist_insert = oa.insert and OAlist_update_by_fun = oa.update_by_fun and OAlist_update_by_fun_gr = oa.update_by_fun_gr and OAlist_filter = oa.filter and OAlist_map2_val_neutr = oa.map2_val_neutr and OAlist_eq = oa.oalist_eq apply standard subgoal by (fact oalist_inv_list_of_oalist) subgoal by (simp only: list_of_oalist_OAlist sort_oalist_ko_def) subgoal by (fact OAlist_list_of_oalist) done global_interpretation oa: oalist_abstract3 "\x. x" "list_of_oalist::('a, 'b) oalist \ ('a, 'b::zero, 'a key_order) oalist_raw" OAlist "list_of_oalist::('a, 'c) oalist \ ('a, 'c::zero, 'a key_order) oalist_raw" OAlist "list_of_oalist::('a, 'd) oalist \ ('a, 'd::zero, 'a key_order) oalist_raw" OAlist defines OAlist_map_val = oa.map_val and OAlist_map2_val = oa.map2_val and OAlist_map2_val_rneutr = oa.map2_val_rneutr and OAlist_lex_ord = oa.lex_ord and OAlist_prod_ord = oa.prod_ord .. lemmas OAlist_lookup_single = oa.lookup_oalist_of_list_single[folded oalist_of_list_def] subsection \Type \oalist_tc\\ text \``tc'' stands for ``type class''.\ global_interpretation tc: comparator "comparator_of" defines lookup_pair_tc = tc.lookup_pair and update_by_pair_tc = tc.update_by_pair and update_by_fun_pair_tc = tc.update_by_fun_pair and update_by_fun_gr_pair_tc = tc.update_by_fun_gr_pair and map2_val_pair_tc = tc.map2_val_pair and lex_ord_pair_tc = tc.lex_ord_pair and prod_ord_pair_tc = tc.prod_ord_pair and sort_oalist_tc = tc.sort_oalist by (fact comparator_of) lemma tc_le_lt [simp]: "tc.le = (\)" "tc.lt = (<)" by (auto simp: le_of_comp_def lt_of_comp_def comparator_of_def intro!: ext split: order.split_asm if_split_asm) typedef (overloaded) ('a, 'b) oalist_tc = "{xs::('a::linorder \ 'b::zero) list. tc.oalist_inv_raw xs}" morphisms list_of_oalist_tc Abs_oalist_tc by (auto intro: tc.oalist_inv_raw_Nil) lemma oalist_tc_eq_iff: "xs = ys \ list_of_oalist_tc xs = list_of_oalist_tc ys" by (simp add: list_of_oalist_tc_inject) lemma oalist_tc_eqI: "list_of_oalist_tc xs = list_of_oalist_tc ys \ xs = ys" by (simp add: oalist_tc_eq_iff) text \Formal, totalized constructor for @{typ "('a, 'b) oalist_tc"}:\ definition OAlist_tc :: "('a \ 'b) list \ ('a::linorder, 'b::zero) oalist_tc" where "OAlist_tc xs = Abs_oalist_tc (sort_oalist_tc xs)" definition "oalist_tc_of_list = OAlist_tc" lemma oalist_inv_list_of_oalist_tc: "tc.oalist_inv_raw (list_of_oalist_tc xs)" using list_of_oalist_tc[of xs] by simp lemma list_of_oalist_OAlist_tc: "list_of_oalist_tc (OAlist_tc xs) = sort_oalist_tc xs" by (simp add: OAlist_tc_def Abs_oalist_tc_inverse tc.oalist_inv_raw_sort_oalist) lemma OAlist_list_of_oalist_tc [code abstype]: "OAlist_tc (list_of_oalist_tc xs) = xs" by (simp add: OAlist_tc_def tc.sort_oalist_id list_of_oalist_tc_inverse oalist_inv_list_of_oalist_tc) lemma list_of_oalist_tc_of_list [code abstract]: "list_of_oalist_tc (oalist_tc_of_list xs) = sort_oalist_tc xs" by (simp add: list_of_oalist_OAlist_tc oalist_tc_of_list_def) lemma list_of_oalist_tc_of_list_id: assumes "tc.oalist_inv_raw xs" shows "list_of_oalist_tc (OAlist_tc xs) = xs" using assms by (simp add: list_of_oalist_OAlist_tc tc.sort_oalist_id) text \It is better to define the following operations directly instead of interpreting @{locale oalist_abstract}, because @{locale oalist_abstract} defines the operations via their \_raw\ analogues, whereas in this case we can define them directly via their \_pair\ analogues.\ definition OAlist_tc_lookup :: "('a::linorder, 'b::zero) oalist_tc \ 'a \ 'b" where "OAlist_tc_lookup xs = lookup_pair_tc (list_of_oalist_tc xs)" definition OAlist_tc_sorted_domain :: "('a::linorder, 'b::zero) oalist_tc \ 'a list" where "OAlist_tc_sorted_domain xs = map fst (list_of_oalist_tc xs)" definition OAlist_tc_empty :: "('a::linorder, 'b::zero) oalist_tc" where "OAlist_tc_empty = OAlist_tc []" definition OAlist_tc_except_min :: "('a, 'b) oalist_tc \ ('a::linorder, 'b::zero) oalist_tc" where "OAlist_tc_except_min xs = OAlist_tc (tl (list_of_oalist_tc xs))" definition OAlist_tc_min_key_val :: "('a::linorder, 'b::zero) oalist_tc \ ('a \ 'b)" where "OAlist_tc_min_key_val xs = hd (list_of_oalist_tc xs)" definition OAlist_tc_insert :: "('a \ 'b) \ ('a, 'b) oalist_tc \ ('a::linorder, 'b::zero) oalist_tc" where "OAlist_tc_insert x xs = OAlist_tc (update_by_pair_tc x (list_of_oalist_tc xs))" definition OAlist_tc_update_by_fun :: "'a \ ('b \ 'b) \ ('a, 'b) oalist_tc \ ('a::linorder, 'b::zero) oalist_tc" where "OAlist_tc_update_by_fun k f xs = OAlist_tc (update_by_fun_pair_tc k f (list_of_oalist_tc xs))" definition OAlist_tc_update_by_fun_gr :: "'a \ ('b \ 'b) \ ('a, 'b) oalist_tc \ ('a::linorder, 'b::zero) oalist_tc" where "OAlist_tc_update_by_fun_gr k f xs = OAlist_tc (update_by_fun_gr_pair_tc k f (list_of_oalist_tc xs))" definition OAlist_tc_filter :: "(('a \ 'b) \ bool) \ ('a, 'b) oalist_tc \ ('a::linorder, 'b::zero) oalist_tc" where "OAlist_tc_filter P xs = OAlist_tc (filter P (list_of_oalist_tc xs))" definition OAlist_tc_map_val :: "('a \ 'b \ 'c) \ ('a, 'b::zero) oalist_tc \ ('a::linorder, 'c::zero) oalist_tc" where "OAlist_tc_map_val f xs = OAlist_tc (map_val_pair f (list_of_oalist_tc xs))" definition OAlist_tc_map2_val :: "('a \ 'b \ 'c \ 'd) \ ('a, 'b::zero) oalist_tc \ ('a, 'c::zero) oalist_tc \ ('a::linorder, 'd::zero) oalist_tc" where "OAlist_tc_map2_val f xs ys = OAlist_tc (map2_val_pair_tc f (map_val_pair (\k b. f k b 0)) (map_val_pair (\k. f k 0)) (list_of_oalist_tc xs) (list_of_oalist_tc ys))" definition OAlist_tc_map2_val_rneutr :: "('a \ 'b \ 'c \ 'b) \ ('a, 'b) oalist_tc \ ('a, 'c::zero) oalist_tc \ ('a::linorder, 'b::zero) oalist_tc" where "OAlist_tc_map2_val_rneutr f xs ys = OAlist_tc (map2_val_pair_tc f id (map_val_pair (\k. f k 0)) (list_of_oalist_tc xs) (list_of_oalist_tc ys))" definition OAlist_tc_map2_val_neutr :: "('a \ 'b \ 'b \ 'b) \ ('a, 'b) oalist_tc \ ('a, 'b) oalist_tc \ ('a::linorder, 'b::zero) oalist_tc" where "OAlist_tc_map2_val_neutr f xs ys = OAlist_tc (map2_val_pair_tc f id id (list_of_oalist_tc xs) (list_of_oalist_tc ys))" definition OAlist_tc_lex_ord :: "('a \ ('b, 'c) comp_opt) \ (('a, 'b::zero) oalist_tc, ('a::linorder, 'c::zero) oalist_tc) comp_opt" where "OAlist_tc_lex_ord f xs ys = lex_ord_pair_tc f (list_of_oalist_tc xs) (list_of_oalist_tc ys)" definition OAlist_tc_prod_ord :: "('a \ 'b \ 'c \ bool) \ ('a, 'b::zero) oalist_tc \ ('a::linorder, 'c::zero) oalist_tc \ bool" where "OAlist_tc_prod_ord f xs ys = prod_ord_pair_tc f (list_of_oalist_tc xs) (list_of_oalist_tc ys)" subsubsection \@{const OAlist_tc_lookup}\ lemma OAlist_tc_lookup_eq_valueI: "(k, v) \ set (list_of_oalist_tc xs) \ OAlist_tc_lookup xs k = v" unfolding OAlist_tc_lookup_def using oalist_inv_list_of_oalist_tc by (rule tc.lookup_pair_eq_valueI) lemma OAlist_tc_lookup_inj: "OAlist_tc_lookup xs = OAlist_tc_lookup ys \ xs = ys" by (simp add: OAlist_tc_lookup_def oalist_inv_list_of_oalist_tc oalist_tc_eqI tc.lookup_pair_inj) lemma OAlist_tc_lookup_oalist_of_list: "distinct (map fst xs) \ OAlist_tc_lookup (oalist_tc_of_list xs) = lookup_dflt xs" by (simp add: OAlist_tc_lookup_def list_of_oalist_tc_of_list tc.lookup_pair_sort_oalist') subsubsection \@{const OAlist_tc_sorted_domain}\ lemma set_OAlist_tc_sorted_domain: "set (OAlist_tc_sorted_domain xs) = fst ` set (list_of_oalist_tc xs)" unfolding OAlist_tc_sorted_domain_def by simp lemma in_OAlist_tc_sorted_domain_iff_lookup: "k \ set (OAlist_tc_sorted_domain xs) \ (OAlist_tc_lookup xs k \ 0)" unfolding OAlist_tc_sorted_domain_def OAlist_tc_lookup_def using oalist_inv_list_of_oalist_tc tc.lookup_pair_eq_0 by fastforce lemma sorted_OAlist_tc_sorted_domain: "sorted_wrt (<) (OAlist_tc_sorted_domain xs)" unfolding OAlist_tc_sorted_domain_def tc_le_lt[symmetric] using oalist_inv_list_of_oalist_tc by (rule tc.oalist_inv_rawD2) subsubsection \@{const OAlist_tc_empty} and Singletons\ lemma list_of_oalist_OAlist_tc_empty [simp, code abstract]: "list_of_oalist_tc OAlist_tc_empty = []" unfolding OAlist_tc_empty_def using tc.oalist_inv_raw_Nil by (rule list_of_oalist_tc_of_list_id) lemma lookup_OAlist_tc_empty: "OAlist_tc_lookup OAlist_tc_empty k = 0" by (simp add: OAlist_tc_lookup_def) lemma OAlist_tc_lookup_single: "OAlist_tc_lookup (oalist_tc_of_list [(k, v)]) k' = (if k = k' then v else 0)" by (simp add: OAlist_tc_lookup_def list_of_oalist_tc_of_list tc.sort_oalist_def comparator_of_def split: order.split) subsubsection \@{const OAlist_tc_except_min}\ lemma list_of_oalist_OAlist_tc_except_min [simp, code abstract]: "list_of_oalist_tc (OAlist_tc_except_min xs) = tl (list_of_oalist_tc xs)" unfolding OAlist_tc_except_min_def by (rule list_of_oalist_tc_of_list_id, rule tc.oalist_inv_raw_tl, fact oalist_inv_list_of_oalist_tc) lemma lookup_OAlist_tc_except_min: "OAlist_tc_lookup (OAlist_tc_except_min xs) k = (if (\k'\fst ` set (list_of_oalist_tc xs). k \ k') then 0 else OAlist_tc_lookup xs k)" by (simp add: OAlist_tc_lookup_def tc.lookup_pair_tl oalist_inv_list_of_oalist_tc split del: if_split cong: if_cong) subsubsection \@{const OAlist_tc_min_key_val}\ lemma OAlist_tc_min_key_val_in: assumes "list_of_oalist_tc xs \ []" shows "OAlist_tc_min_key_val xs \ set (list_of_oalist_tc xs)" unfolding OAlist_tc_min_key_val_def using assms by simp lemma snd_OAlist_tc_min_key_val: assumes "list_of_oalist_tc xs \ []" shows "snd (OAlist_tc_min_key_val xs) = OAlist_tc_lookup xs (fst (OAlist_tc_min_key_val xs))" proof - let ?xs = "list_of_oalist_tc xs" from assms have *: "OAlist_tc_min_key_val xs \ set ?xs" by (rule OAlist_tc_min_key_val_in) show ?thesis unfolding OAlist_tc_lookup_def by (rule HOL.sym, rule tc.lookup_pair_eq_valueI, fact oalist_inv_list_of_oalist_tc, simp add: *) qed lemma OAlist_tc_min_key_val_minimal: assumes "z \ set (list_of_oalist_tc xs)" shows "fst (OAlist_tc_min_key_val xs) \ fst z" proof - let ?xs = "list_of_oalist_tc xs" from assms have "?xs \ []" by auto hence "OAlist_tc_sorted_domain xs \ []" by (simp add: OAlist_tc_sorted_domain_def) then obtain h xs' where eq: "OAlist_tc_sorted_domain xs = h # xs'" using list.exhaust by blast with sorted_OAlist_tc_sorted_domain[of xs] have *: "\y\set xs'. h < y" by simp from assms have "fst z \ set (OAlist_tc_sorted_domain xs)" by (simp add: OAlist_tc_sorted_domain_def) hence disj: "fst z = h \ fst z \ set xs'" by (simp add: eq) from \?xs \ []\ have "fst (OAlist_tc_min_key_val xs) = hd (OAlist_tc_sorted_domain xs)" by (simp add: OAlist_tc_min_key_val_def OAlist_tc_sorted_domain_def hd_map) also have "... = h" by (simp add: eq) also from disj have "... \ fst z" proof assume "fst z = h" thus ?thesis by simp next assume "fst z \ set xs'" with * have "h < fst z" .. thus ?thesis by simp qed finally show ?thesis . qed subsubsection \@{const OAlist_tc_insert}\ lemma list_of_oalist_OAlist_tc_insert [simp, code abstract]: "list_of_oalist_tc (OAlist_tc_insert x xs) = update_by_pair_tc x (list_of_oalist_tc xs)" unfolding OAlist_tc_insert_def by (rule list_of_oalist_tc_of_list_id, rule tc.oalist_inv_raw_update_by_pair, fact oalist_inv_list_of_oalist_tc) lemma lookup_OAlist_tc_insert: "OAlist_tc_lookup (OAlist_tc_insert (k, v) xs) k' = (if k = k' then v else OAlist_tc_lookup xs k')" by (simp add: OAlist_tc_lookup_def tc.lookup_pair_update_by_pair oalist_inv_list_of_oalist_tc split del: if_split cong: if_cong) subsubsection \@{const OAlist_tc_update_by_fun} and @{const OAlist_tc_update_by_fun_gr}\ lemma list_of_oalist_OAlist_tc_update_by_fun [simp, code abstract]: "list_of_oalist_tc (OAlist_tc_update_by_fun k f xs) = update_by_fun_pair_tc k f (list_of_oalist_tc xs)" unfolding OAlist_tc_update_by_fun_def by (rule list_of_oalist_tc_of_list_id, rule tc.oalist_inv_raw_update_by_fun_pair, fact oalist_inv_list_of_oalist_tc) lemma lookup_OAlist_tc_update_by_fun: "OAlist_tc_lookup (OAlist_tc_update_by_fun k f xs) k' = (if k = k' then f else id) (OAlist_tc_lookup xs k')" by (simp add: OAlist_tc_lookup_def tc.lookup_pair_update_by_fun_pair oalist_inv_list_of_oalist_tc split del: if_split cong: if_cong) lemma list_of_oalist_OAlist_tc_update_by_fun_gr [simp, code abstract]: "list_of_oalist_tc (OAlist_tc_update_by_fun_gr k f xs) = update_by_fun_gr_pair_tc k f (list_of_oalist_tc xs)" unfolding OAlist_tc_update_by_fun_gr_def by (rule list_of_oalist_tc_of_list_id, rule tc.oalist_inv_raw_update_by_fun_gr_pair, fact oalist_inv_list_of_oalist_tc) lemma OAlist_tc_update_by_fun_gr_eq_OAlist_tc_update_by_fun: "OAlist_tc_update_by_fun_gr = OAlist_tc_update_by_fun" by (rule, rule, rule, simp add: OAlist_tc_update_by_fun_gr_def OAlist_tc_update_by_fun_def tc.update_by_fun_gr_pair_eq_update_by_fun_pair oalist_inv_list_of_oalist_tc) subsubsection \@{const OAlist_tc_filter}\ lemma list_of_oalist_OAlist_tc_filter [simp, code abstract]: "list_of_oalist_tc (OAlist_tc_filter P xs) = filter P (list_of_oalist_tc xs)" unfolding OAlist_tc_filter_def by (rule list_of_oalist_tc_of_list_id, rule tc.oalist_inv_raw_filter, fact oalist_inv_list_of_oalist_tc) lemma lookup_OAlist_tc_filter: "OAlist_tc_lookup (OAlist_tc_filter P xs) k = (let v = OAlist_tc_lookup xs k in if P (k, v) then v else 0)" by (simp add: OAlist_tc_lookup_def tc.lookup_pair_filter oalist_inv_list_of_oalist_tc) subsubsection \@{const OAlist_tc_map_val}\ lemma list_of_oalist_OAlist_tc_map_val [simp, code abstract]: "list_of_oalist_tc (OAlist_tc_map_val f xs) = map_val_pair f (list_of_oalist_tc xs)" unfolding OAlist_tc_map_val_def by (rule list_of_oalist_tc_of_list_id, rule tc.oalist_inv_raw_map_val_pair, fact oalist_inv_list_of_oalist_tc) lemma OAlist_tc_map_val_cong: assumes "\k v. (k, v) \ set (list_of_oalist_tc xs) \ f k v = g k v" shows "OAlist_tc_map_val f xs = OAlist_tc_map_val g xs" unfolding OAlist_tc_map_val_def by (rule arg_cong[where f=OAlist_tc], rule tc.map_val_pair_cong, elim assms) lemma lookup_OAlist_tc_map_val: "f k 0 = 0 \ OAlist_tc_lookup (OAlist_tc_map_val f xs) k = f k (OAlist_tc_lookup xs k)" by (simp add: OAlist_tc_lookup_def tc.lookup_pair_map_val_pair oalist_inv_list_of_oalist_tc) subsubsection \@{const OAlist_tc_map2_val} @{const OAlist_tc_map2_val_rneutr} and @{const OAlist_tc_map2_val_neutr}\ lemma list_of_oalist_map2_val [simp, code abstract]: "list_of_oalist_tc (OAlist_tc_map2_val f xs ys) = map2_val_pair_tc f (map_val_pair (\k b. f k b 0)) (map_val_pair (\k. f k 0)) (list_of_oalist_tc xs) (list_of_oalist_tc ys)" unfolding OAlist_tc_map2_val_def by (rule list_of_oalist_tc_of_list_id, rule tc.oalist_inv_raw_map2_val_pair, fact oalist_inv_list_of_oalist_tc, fact oalist_inv_list_of_oalist_tc, fact tc.map2_val_compat_map_val_pair, fact tc.map2_val_compat_map_val_pair) lemma list_of_oalist_OAlist_tc_map2_val_rneutr [simp, code abstract]: "list_of_oalist_tc (OAlist_tc_map2_val_rneutr f xs ys) = map2_val_pair_tc f id (map_val_pair (\k c. f k 0 c)) (list_of_oalist_tc xs) (list_of_oalist_tc ys)" unfolding OAlist_tc_map2_val_rneutr_def by (rule list_of_oalist_tc_of_list_id, rule tc.oalist_inv_raw_map2_val_pair, fact oalist_inv_list_of_oalist_tc, fact oalist_inv_list_of_oalist_tc, fact tc.map2_val_compat_id, fact tc.map2_val_compat_map_val_pair) lemma list_of_oalist_OAlist_tc_map2_val_neutr [simp, code abstract]: "list_of_oalist_tc (OAlist_tc_map2_val_neutr f xs ys) = map2_val_pair_tc f id id (list_of_oalist_tc xs) (list_of_oalist_tc ys)" unfolding OAlist_tc_map2_val_neutr_def by (rule list_of_oalist_tc_of_list_id, rule tc.oalist_inv_raw_map2_val_pair, fact oalist_inv_list_of_oalist_tc, fact oalist_inv_list_of_oalist_tc, fact tc.map2_val_compat_id, fact tc.map2_val_compat_id) lemma lookup_OAlist_tc_map2_val: assumes "\k. f k 0 0 = 0" shows "OAlist_tc_lookup (OAlist_tc_map2_val f xs ys) k = f k (OAlist_tc_lookup xs k) (OAlist_tc_lookup ys k)" by (simp add: OAlist_tc_lookup_def tc.lookup_pair_map2_val_pair tc.map2_val_compat_map_val_pair assms oalist_inv_list_of_oalist_tc) lemma lookup_OAlist_tc_map2_val_rneutr: assumes "\k x. f k x 0 = x" shows "OAlist_tc_lookup (OAlist_tc_map2_val_rneutr f xs ys) k = f k (OAlist_tc_lookup xs k) (OAlist_tc_lookup ys k)" proof (simp add: OAlist_tc_lookup_def, rule tc.lookup_pair_map2_val_pair) fix zs::"('a \ 'b) list" assume "tc.oalist_inv_raw zs" thus "id zs = map_val_pair (\k v. f k v 0) zs" by (simp add: assms tc.map_pair_id) qed (fact oalist_inv_list_of_oalist_tc, fact oalist_inv_list_of_oalist_tc, fact tc.map2_val_compat_id, fact tc.map2_val_compat_map_val_pair, rule refl, simp only: assms) lemma lookup_OAlist_tc_map2_val_neutr: assumes "\k x. f k x 0 = x" and "\k x. f k 0 x = x" shows "OAlist_tc_lookup (OAlist_tc_map2_val_neutr f xs ys) k = f k (OAlist_tc_lookup xs k) (OAlist_tc_lookup ys k)" proof (simp add: OAlist_tc_lookup_def, rule tc.lookup_pair_map2_val_pair) fix zs::"('a \ 'b) list" assume "tc.oalist_inv_raw zs" thus "id zs = map_val_pair (\k v. f k v 0) zs" by (simp add: assms(1) tc.map_pair_id) next fix zs::"('a \ 'b) list" assume "tc.oalist_inv_raw zs" thus "id zs = map_val_pair (\k. f k 0) zs" by (simp add: assms(2) tc.map_pair_id) qed (fact oalist_inv_list_of_oalist_tc, fact oalist_inv_list_of_oalist_tc, fact tc.map2_val_compat_id, fact tc.map2_val_compat_id, simp only: assms(1)) lemma OAlist_tc_map2_val_rneutr_singleton_eq_OAlist_tc_update_by_fun: assumes "\a x. f a x 0 = x" and "list_of_oalist_tc ys = [(k, v)]" shows "OAlist_tc_map2_val_rneutr f xs ys = OAlist_tc_update_by_fun k (\x. f k x v) xs" by (simp add: OAlist_tc_map2_val_rneutr_def OAlist_tc_update_by_fun_def assms tc.map2_val_pair_singleton_eq_update_by_fun_pair oalist_inv_list_of_oalist_tc) subsubsection \@{const OAlist_tc_lex_ord} and @{const OAlist_tc_prod_ord}\ lemma OAlist_tc_lex_ord_EqI: "(\k. k \ fst ` set (list_of_oalist_tc xs) \ fst ` set (list_of_oalist_tc ys) \ f k (OAlist_tc_lookup xs k) (OAlist_tc_lookup ys k) = Some Eq) \ OAlist_tc_lex_ord f xs ys = Some Eq" by (simp add: OAlist_tc_lex_ord_def OAlist_tc_lookup_def, rule tc.lex_ord_pair_EqI, rule oalist_inv_list_of_oalist_tc, rule oalist_inv_list_of_oalist_tc, blast) lemma OAlist_tc_lex_ord_valI: assumes "aux \ Some Eq" and "k \ fst ` set (list_of_oalist_tc xs) \ fst ` set (list_of_oalist_tc ys)" shows "aux = f k (OAlist_tc_lookup xs k) (OAlist_tc_lookup ys k) \ (\k'. k' \ fst ` set (list_of_oalist_tc xs) \ fst ` set (list_of_oalist_tc ys) \ k' < k \ f k' (OAlist_tc_lookup xs k') (OAlist_tc_lookup ys k') = Some Eq) \ OAlist_tc_lex_ord f xs ys = aux" by (simp (no_asm_use) add: OAlist_tc_lex_ord_def OAlist_tc_lookup_def, rule tc.lex_ord_pair_valI, rule oalist_inv_list_of_oalist_tc, rule oalist_inv_list_of_oalist_tc, rule assms(1), rule assms(2), simp_all) lemma OAlist_tc_lex_ord_EqD: "OAlist_tc_lex_ord f xs ys = Some Eq \ k \ fst ` set (list_of_oalist_tc xs) \ fst ` set (list_of_oalist_tc ys) \ f k (OAlist_tc_lookup xs k) (OAlist_tc_lookup ys k) = Some Eq" by (simp add: OAlist_tc_lex_ord_def OAlist_tc_lookup_def, rule tc.lex_ord_pair_EqD[where f=f], rule oalist_inv_list_of_oalist_tc, rule oalist_inv_list_of_oalist_tc, assumption, simp) lemma OAlist_tc_lex_ord_valE: assumes "OAlist_tc_lex_ord f xs ys = aux" and "aux \ Some Eq" obtains k where "k \ fst ` set (list_of_oalist_tc xs) \ fst ` set (list_of_oalist_tc ys)" and "aux = f k (OAlist_tc_lookup xs k) (OAlist_tc_lookup ys k)" and "\k'. k' \ fst ` set (list_of_oalist_tc xs) \ fst ` set (list_of_oalist_tc ys) \ k' < k \ f k' (OAlist_tc_lookup xs k') (OAlist_tc_lookup ys k') = Some Eq" proof - note oalist_inv_list_of_oalist_tc oalist_inv_list_of_oalist_tc moreover from assms(1) have "lex_ord_pair_tc f (list_of_oalist_tc xs) (list_of_oalist_tc ys) = aux" by (simp only: OAlist_tc_lex_ord_def) ultimately obtain k where 1: "k \ fst ` set (list_of_oalist_tc xs) \ fst ` set (list_of_oalist_tc ys)" and "aux = f k (lookup_pair_tc (list_of_oalist_tc xs) k) (lookup_pair_tc (list_of_oalist_tc ys) k)" and "\k'. k' \ fst ` set (list_of_oalist_tc xs) \ fst ` set (list_of_oalist_tc ys) \ k' < k \ f k' (lookup_pair_tc (list_of_oalist_tc xs) k') (lookup_pair_tc (list_of_oalist_tc ys) k') = Some Eq" using assms(2) unfolding tc_le_lt[symmetric] by (rule tc.lex_ord_pair_valE, blast) from this(2, 3) have "aux = f k (OAlist_tc_lookup xs k) (OAlist_tc_lookup ys k)" and "\k'. k' \ fst ` set (list_of_oalist_tc xs) \ fst ` set (list_of_oalist_tc ys) \ k' < k \ f k' (OAlist_tc_lookup xs k') (OAlist_tc_lookup ys k') = Some Eq" by (simp_all only: OAlist_tc_lookup_def) with 1 show ?thesis .. qed lemma OAlist_tc_prod_ord_alt: "OAlist_tc_prod_ord P xs ys \ (\k\fst ` set (list_of_oalist_tc xs) \ fst ` set (list_of_oalist_tc ys). P k (OAlist_tc_lookup xs k) (OAlist_tc_lookup ys k))" by (simp add: OAlist_tc_prod_ord_def OAlist_tc_lookup_def tc.prod_ord_pair_alt oalist_inv_list_of_oalist_tc) subsubsection \Instance of @{class equal}\ instantiation oalist_tc :: (linorder, zero) equal begin definition equal_oalist_tc :: "('a, 'b) oalist_tc \ ('a, 'b) oalist_tc \ bool" where "equal_oalist_tc xs ys = (list_of_oalist_tc xs = list_of_oalist_tc ys)" instance by (intro_classes, simp add: equal_oalist_tc_def list_of_oalist_tc_inject) end subsection \Experiment\ lemma "oalist_tc_of_list [(0::nat, 4::nat), (1, 3), (0, 2), (1, 1)] = oalist_tc_of_list [(0, 4), (1, 3)]" by eval lemma "OAlist_tc_except_min (oalist_tc_of_list ([(1, 3), (0::nat, 4::nat), (0, 2), (1, 1)])) = oalist_tc_of_list [(1, 3)]" by eval lemma "OAlist_tc_min_key_val (oalist_tc_of_list [(1, 3), (0::nat, 4::nat), (0, 2), (1, 1)]) = (0, 4)" by eval lemma "OAlist_tc_lookup (oalist_tc_of_list [(0::nat, 4::nat), (1, 3), (0, 2), (1, 1)]) 1 = 3" by eval lemma "OAlist_tc_prod_ord (\_. greater_eq) (oalist_tc_of_list [(1, 4), (0::nat, 4::nat), (1, 3), (0, 2), (3, 1)]) (oalist_tc_of_list [(0, 4), (1, 3), (2, 2), (1, 1)]) = False" by eval lemma "OAlist_tc_map2_val_rneutr (\_. minus) (oalist_tc_of_list [(1, 4), (0::nat, 4::int), (1, 3), (0, 2), (3, 1)]) (oalist_tc_of_list [(0, 4), (1, 3), (2, 2), (1, 1)]) = oalist_tc_of_list [(1, 1), (2, - 2), (3, 1)]" by eval end (* theory *) diff --git a/thys/Polynomials/Power_Products.thy b/thys/Polynomials/Power_Products.thy --- a/thys/Polynomials/Power_Products.thy +++ b/thys/Polynomials/Power_Products.thy @@ -1,2698 +1,2698 @@ (* Author: Fabian Immler, Alexander Maletzky *) section \Abstract Power-Products\ theory Power_Products imports Complex_Main "HOL-Library.Function_Algebras" "HOL-Library.Countable" "More_MPoly_Type" "Utils" Well_Quasi_Orders.Well_Quasi_Orders begin text \This theory formalizes the concept of "power-products". A power-product can be thought of as the product of some indeterminates, such as $x$, $x^2\,y$, $x\,y^3\,z^7$, etc., without any scalar coefficient. The approach in this theory is to capture the notion of "power-product" (also called "monomial") as type class. A canonical instance for power-product is the type @{typ "'var \\<^sub>0 nat"}, which is interpreted as mapping from variables in the power-product to exponents. A slightly unintuitive (but fitting better with the standard type class instantiations of @{typ "'a \\<^sub>0 'b"}) approach is to write addition to denote "multiplication" of power products. For example, $x^2y$ would be represented as a function \p = (X \ 2, Y \ 1)\, $xz$ as a function \q = (X \ 1, Z \ 1)\. With the (pointwise) instantiation of addition of @{typ "'a \\<^sub>0 'b"}, we will write \p + q = (X \ 3, Y \ 1, Z \ 1)\ for the product $x^2y \cdot xz = x^3yz$ \ subsection \Constant @{term Keys}\ text \Legacy:\ lemmas keys_eq_empty_iff = keys_eq_empty definition Keys :: "('a \\<^sub>0 'b::zero) set \ 'a set" where "Keys F = \(keys ` F)" lemma in_Keys: "s \ Keys F \ (\f\F. s \ keys f)" unfolding Keys_def by simp lemma in_KeysI: assumes "s \ keys f" and "f \ F" shows "s \ Keys F" unfolding in_Keys using assms .. lemma in_KeysE: assumes "s \ Keys F" obtains f where "s \ keys f" and "f \ F" using assms unfolding in_Keys .. lemma Keys_mono: assumes "A \ B" shows "Keys A \ Keys B" using assms by (auto simp add: Keys_def) lemma Keys_insert: "Keys (insert a A) = keys a \ Keys A" by (simp add: Keys_def) lemma Keys_Un: "Keys (A \ B) = Keys A \ Keys B" by (simp add: Keys_def) lemma finite_Keys: assumes "finite A" shows "finite (Keys A)" unfolding Keys_def by (rule, fact assms, rule finite_keys) lemma Keys_not_empty: assumes "a \ A" and "a \ 0" shows "Keys A \ {}" proof assume "Keys A = {}" from \a \ 0\ have "keys a \ {}" using aux by fastforce then obtain s where "s \ keys a" by blast from this assms(1) have "s \ Keys A" by (rule in_KeysI) with \Keys A = {}\ show False by simp qed lemma Keys_empty [simp]: "Keys {} = {}" by (simp add: Keys_def) lemma Keys_zero [simp]: "Keys {0} = {}" by (simp add: Keys_def) lemma keys_subset_Keys: assumes "f \ F" shows "keys f \ Keys F" using in_KeysI[OF _ assms] by auto lemma Keys_minus: "Keys (A - B) \ Keys A" by (auto simp add: Keys_def) lemma Keys_minus_zero: "Keys (A - {0}) = Keys A" proof (cases "0 \ A") case True hence "(A - {0}) \ {0} = A" by auto hence "Keys A = Keys ((A - {0}) \ {0})" by simp also have "... = Keys (A - {0}) \ Keys {0::('a \\<^sub>0 'b)}" by (fact Keys_Un) also have "... = Keys (A - {0})" by simp finally show ?thesis by simp next case False hence "A - {0} = A" by simp thus ?thesis by simp qed subsection \Constant @{term except}\ definition except_fun :: "('a \ 'b) \ 'a set \ ('a \ 'b::zero)" where "except_fun f S = (\x. (f x when x \ S))" lift_definition except :: "('a \\<^sub>0 'b) \ 'a set \ ('a \\<^sub>0 'b::zero)" is except_fun proof - fix p::"'a \ 'b" and S::"'a set" assume "finite {t. p t \ 0}" show "finite {t. except_fun p S t \ 0}" proof (rule finite_subset[of _ "{t. p t \ 0}"], rule) fix u assume "u \ {t. except_fun p S t \ 0}" hence "p u \ 0" by (simp add: except_fun_def) thus "u \ {t. p t \ 0}" by simp qed fact qed lemma lookup_except_when: "lookup (except p S) = (\t. lookup p t when t \ S)" by (auto simp: except.rep_eq except_fun_def) lemma lookup_except: "lookup (except p S) = (\t. if t \ S then 0 else lookup p t)" by (rule ext) (simp add: lookup_except_when) lemma lookup_except_singleton: "lookup (except p {t}) t = 0" by (simp add: lookup_except) lemma except_zero [simp]: "except 0 S = 0" by (rule poly_mapping_eqI) (simp add: lookup_except) lemma lookup_except_eq_idI: assumes "t \ S" shows "lookup (except p S) t = lookup p t" using assms by (simp add: lookup_except) lemma lookup_except_eq_zeroI: assumes "t \ S" shows "lookup (except p S) t = 0" using assms by (simp add: lookup_except) lemma except_empty [simp]: "except p {} = p" by (rule poly_mapping_eqI) (simp add: lookup_except) lemma except_eq_zeroI: assumes "keys p \ S" shows "except p S = 0" proof (rule poly_mapping_eqI, simp) fix t show "lookup (except p S) t = 0" proof (cases "t \ S") case True thus ?thesis by (rule lookup_except_eq_zeroI) next case False then show ?thesis by (metis assms in_keys_iff lookup_except_eq_idI subset_eq) qed qed lemma except_eq_zeroE: assumes "except p S = 0" shows "keys p \ S" by (metis assms aux in_keys_iff lookup_except_eq_idI subset_iff) lemma except_eq_zero_iff: "except p S = 0 \ keys p \ S" by (rule, elim except_eq_zeroE, elim except_eq_zeroI) lemma except_keys [simp]: "except p (keys p) = 0" by (rule except_eq_zeroI, rule subset_refl) lemma plus_except: "p = Poly_Mapping.single t (lookup p t) + except p {t}" by (rule poly_mapping_eqI, simp add: lookup_add lookup_single lookup_except when_def split: if_split) lemma keys_except: "keys (except p S) = keys p - S" by (transfer, auto simp: except_fun_def) lemma except_single: "except (Poly_Mapping.single u c) S = (Poly_Mapping.single u c when u \ S)" by (rule poly_mapping_eqI) (simp add: lookup_except lookup_single when_def) lemma except_plus: "except (p + q) S = except p S + except q S" by (rule poly_mapping_eqI) (simp add: lookup_except lookup_add) lemma except_minus: "except (p - q) S = except p S - except q S" by (rule poly_mapping_eqI) (simp add: lookup_except lookup_minus) lemma except_uminus: "except (- p) S = - except p S" by (rule poly_mapping_eqI) (simp add: lookup_except) lemma except_except: "except (except p S) T = except p (S \ T)" by (rule poly_mapping_eqI) (simp add: lookup_except) lemma poly_mapping_keys_eqI: assumes a1: "keys p = keys q" and a2: "\t. t \ keys p \ lookup p t = lookup q t" shows "p = q" proof (rule poly_mapping_eqI) fix t show "lookup p t = lookup q t" proof (cases "t \ keys p") case True thus ?thesis by (rule a2) next case False moreover from this have "t \ keys q" unfolding a1 . ultimately have "lookup p t = 0" and "lookup q t = 0" unfolding in_keys_iff by simp_all thus ?thesis by simp qed qed lemma except_id_iff: "except p S = p \ keys p \ S = {}" by (metis Diff_Diff_Int Diff_eq_empty_iff Diff_triv inf_le2 keys_except lookup_except_eq_idI lookup_except_eq_zeroI not_in_keys_iff_lookup_eq_zero poly_mapping_keys_eqI) lemma keys_subset_wf: "wfP (\p q::('a, 'b::zero) poly_mapping. keys p \ keys q)" unfolding wfP_def proof (intro wfI_min) fix x::"('a, 'b) poly_mapping" and Q assume x_in: "x \ Q" let ?Q0 = "card ` keys ` Q" from x_in have "card (keys x) \ ?Q0" by simp from wfE_min[OF wf this] obtain z0 where z0_in: "z0 \ ?Q0" and z0_min: "\y. (y, z0) \ {(x, y). x < y} \ y \ ?Q0" by auto from z0_in obtain z where z0_def: "z0 = card (keys z)" and "z \ Q" by auto show "\z\Q. \y. (y, z) \ {(p, q). keys p \ keys q} \ y \ Q" proof (intro bexI[of _ z], rule, rule) fix y::"('a, 'b) poly_mapping" let ?y0 = "card (keys y)" assume "(y, z) \ {(p, q). keys p \ keys q}" hence "keys y \ keys z" by simp hence "?y0 < z0" unfolding z0_def by (simp add: psubset_card_mono) hence "(?y0, z0) \ {(x, y). x < y}" by simp from z0_min[OF this] show "y \ Q" by auto qed (fact) qed lemma poly_mapping_except_induct: assumes base: "P 0" and ind: "\p t. p \ 0 \ t \ keys p \ P (except p {t}) \ P p" shows "P p" proof (induct rule: wfP_induct[OF keys_subset_wf]) fix p::"('a, 'b) poly_mapping" assume "\q. keys q \ keys p \ P q" hence IH: "\q. keys q \ keys p \ P q" by simp show "P p" proof (cases "p = 0") case True thus ?thesis using base by simp next case False hence "keys p \ {}" by simp then obtain t where "t \ keys p" by blast show ?thesis proof (rule ind, fact, fact, rule IH, simp only: keys_except, rule, rule Diff_subset, rule) assume "keys p - {t} = keys p" hence "t \ keys p" by blast from this \t \ keys p\ show False .. qed qed qed lemma poly_mapping_except_induct': assumes "\p. (\t. t \ keys p \ P (except p {t})) \ P p" shows "P p" proof (induct "card (keys p)" arbitrary: p) case 0 with finite_keys[of p] have "keys p = {}" by simp show ?case by (rule assms, simp add: \keys p = {}\) next case step: (Suc n) show ?case proof (rule assms) fix t assume "t \ keys p" show "P (except p {t})" proof (rule step(1), simp add: keys_except) from step(2) \t \ keys p\ finite_keys[of p] show "n = card (keys p - {t})" by simp qed qed qed lemma poly_mapping_plus_induct: assumes "P 0" and "\p c t. c \ 0 \ t \ keys p \ P p \ P (Poly_Mapping.single t c + p)" shows "P p" proof (induct "card (keys p)" arbitrary: p) case 0 with finite_keys[of p] have "keys p = {}" by simp hence "p = 0" by simp with assms(1) show ?case by simp next case step: (Suc n) from step(2) obtain t where t: "t \ keys p" by (metis card_eq_SucD insert_iff) define c where "c = lookup p t" define q where "q = except p {t}" have *: "p = Poly_Mapping.single t c + q" by (rule poly_mapping_eqI, simp add: lookup_add lookup_single Poly_Mapping.when_def, intro conjI impI, simp add: q_def lookup_except c_def, simp add: q_def lookup_except_eq_idI) show ?case proof (simp only: *, rule assms(2)) from t show "c \ 0" using c_def by auto next show "t \ keys q" by (simp add: q_def keys_except) next show "P q" proof (rule step(1)) from step(2) \t \ keys p\ show "n = card (keys q)" unfolding q_def keys_except by (metis Suc_inject card.remove finite_keys) qed qed qed lemma except_Diff_singleton: "except p (keys p - {t}) = Poly_Mapping.single t (lookup p t)" by (rule poly_mapping_eqI) (simp add: lookup_single in_keys_iff lookup_except when_def) lemma except_Un_plus_Int: "except p (U \ V) + except p (U \ V) = except p U + except p V" by (rule poly_mapping_eqI) (simp add: lookup_except lookup_add) corollary except_Int: assumes "keys p \ U \ V" shows "except p (U \ V) = except p U + except p V" proof - from assms have "except p (U \ V) = 0" by (rule except_eq_zeroI) hence "except p (U \ V) = except p (U \ V) + except p (U \ V)" by simp also have "\ = except p U + except p V" by (fact except_Un_plus_Int) finally show ?thesis . qed lemma except_keys_Int [simp]: "except p (keys p \ U) = except p U" by (rule poly_mapping_eqI) (simp add: in_keys_iff lookup_except) lemma except_Int_keys [simp]: "except p (U \ keys p) = except p U" by (simp only: Int_commute[of U] except_keys_Int) lemma except_keys_Diff: "except p (keys p - U) = except p (- U)" proof - have "except p (keys p - U) = except p (keys p \ (- U))" by (simp only: Diff_eq) also have "\ = except p (- U)" by simp finally show ?thesis . qed lemma except_decomp: "p = except p U + except p (- U)" by (rule poly_mapping_eqI) (simp add: lookup_except lookup_add) corollary except_Compl: "except p (- U) = p - except p U" by (metis add_diff_cancel_left' except_decomp) subsection \'Divisibility' on Additive Structures\ context plus begin definition adds :: "'a \ 'a \ bool" (infix "adds" 50) where "b adds a \ (\k. a = b + k)" lemma addsI [intro?]: "a = b + k \ b adds a" unfolding adds_def .. lemma addsE [elim?]: "b adds a \ (\k. a = b + k \ P) \ P" unfolding adds_def by blast end context comm_monoid_add begin lemma adds_refl [simp]: "a adds a" proof show "a = a + 0" by simp qed lemma adds_trans [trans]: assumes "a adds b" and "b adds c" shows "a adds c" proof - from assms obtain v where "b = a + v" by (auto elim!: addsE) moreover from assms obtain w where "c = b + w" by (auto elim!: addsE) ultimately have "c = a + (v + w)" by (simp add: add.assoc) then show ?thesis .. qed lemma subset_divisors_adds: "{c. c adds a} \ {c. c adds b} \ a adds b" by (auto simp add: subset_iff intro: adds_trans) lemma strict_subset_divisors_adds: "{c. c adds a} \ {c. c adds b} \ a adds b \ \ b adds a" by (auto simp add: subset_iff intro: adds_trans) lemma zero_adds [simp]: "0 adds a" by (auto intro!: addsI) lemma adds_plus_right [simp]: "a adds c \ a adds (b + c)" by (auto intro!: add.left_commute addsI elim!: addsE) lemma adds_plus_left [simp]: "a adds b \ a adds (b + c)" using adds_plus_right [of a b c] by (simp add: ac_simps) lemma adds_triv_right [simp]: "a adds b + a" by (rule adds_plus_right) (rule adds_refl) lemma adds_triv_left [simp]: "a adds a + b" by (rule adds_plus_left) (rule adds_refl) lemma plus_adds_mono: assumes "a adds b" and "c adds d" shows "a + c adds b + d" proof - from \a adds b\ obtain b' where "b = a + b'" .. moreover from \c adds d\ obtain d' where "d = c + d'" .. ultimately have "b + d = (a + c) + (b' + d')" by (simp add: ac_simps) then show ?thesis .. qed lemma plus_adds_left: "a + b adds c \ a adds c" by (simp add: adds_def add.assoc) blast lemma plus_adds_right: "a + b adds c \ b adds c" using plus_adds_left [of b a c] by (simp add: ac_simps) end class ninv_comm_monoid_add = comm_monoid_add + assumes plus_eq_zero: "s + t = 0 \ s = 0" begin lemma plus_eq_zero_2: "t = 0" if "s + t = 0" using that by (simp only: add_commute[of s t] plus_eq_zero) lemma adds_zero: "s adds 0 \ (s = 0)" proof assume "s adds 0" from this obtain k where "0 = s + k" unfolding adds_def .. from this plus_eq_zero[of s k] show "s = 0" by blast next assume "s = 0" thus "s adds 0" by simp qed end context canonically_ordered_monoid_add begin subclass ninv_comm_monoid_add by (standard, simp) end class comm_powerprod = cancel_comm_monoid_add begin lemma adds_canc: "s + u adds t + u \ s adds t" for s t u::'a unfolding adds_def apply auto apply (metis local.add.left_commute local.add_diff_cancel_left' local.add_diff_cancel_right') using add_assoc add_commute by auto lemma adds_canc_2: "u + s adds u + t \ s adds t" by (simp add: adds_canc ac_simps) lemma add_minus_2: "(s + t) - s = t" by simp lemma adds_minus: assumes "s adds t" shows "(t - s) + s = t" proof - from assms adds_def[of s t] obtain u where u: "t = u + s" by (auto simp: ac_simps) then have "t - s = u" by simp thus ?thesis using u by simp qed lemma plus_adds_0: assumes "(s + t) adds u" shows "s adds (u - t)" proof - from assms have "(s + t) adds ((u - t) + t)" using adds_minus local.plus_adds_right by presburger thus ?thesis using adds_canc[of s t "u - t"] by simp qed lemma plus_adds_2: assumes "t adds u" and "s adds (u - t)" shows "(s + t) adds u" by (metis adds_canc adds_minus assms) lemma plus_adds: shows "(s + t) adds u \ (t adds u \ s adds (u - t))" proof assume a1: "(s + t) adds u" show "t adds u \ s adds (u - t)" proof from plus_adds_right[OF a1] show "t adds u" . next from plus_adds_0[OF a1] show "s adds (u - t)" . qed next assume "t adds u \ s adds (u - t)" hence "t adds u" and "s adds (u - t)" by auto from plus_adds_2[OF \t adds u\ \s adds (u - t)\] show "(s + t) adds u" . qed lemma minus_plus: assumes "s adds t" shows "(t - s) + u = (t + u) - s" proof - from assms obtain k where k: "t = s + k" unfolding adds_def .. hence "t - s = k" by simp also from k have "(t + u) - s = k + u" by (simp add: add_assoc) finally show ?thesis by simp qed lemma minus_plus_minus: assumes "s adds t" and "u adds v" shows "(t - s) + (v - u) = (t + v) - (s + u)" using add_commute assms(1) assms(2) diff_diff_add minus_plus by auto lemma minus_plus_minus_cancel: assumes "u adds t" and "s adds u" shows "(t - u) + (u - s) = t - s" by (metis assms(1) assms(2) local.add_diff_cancel_left' local.add_diff_cancel_right local.addsE minus_plus) end text \Instances of class \lcs_powerprod\ are types of commutative power-products admitting (not necessarily unique) least common sums (inspired from least common multiplies). Note that if the components of indeterminates are arbitrary integers (as for instance in Laurent polynomials), then no unique lcss exist.\ class lcs_powerprod = comm_powerprod + fixes lcs::"'a \ 'a \ 'a" assumes adds_lcs: "s adds (lcs s t)" assumes lcs_adds: "s adds u \ t adds u \ (lcs s t) adds u" assumes lcs_comm: "lcs s t = lcs t s" begin lemma adds_lcs_2: "t adds (lcs s t)" by (simp only: lcs_comm[of s t], rule adds_lcs) lemma lcs_adds_plus: "lcs s t adds s + t" by (simp add: lcs_adds) text \"gcs" stands for "greatest common summand".\ definition gcs :: "'a \ 'a \ 'a" where "gcs s t = (s + t) - (lcs s t)" lemma gcs_plus_lcs: "(gcs s t) + (lcs s t) = s + t" unfolding gcs_def by (rule adds_minus, fact lcs_adds_plus) lemma gcs_adds: "(gcs s t) adds s" proof - have "t adds (lcs s t)" (is "t adds ?l") unfolding lcs_comm[of s t] by (fact adds_lcs) then obtain u where eq1: "?l = t + u" unfolding adds_def .. from lcs_adds_plus[of s t] obtain v where eq2: "s + t = ?l + v" unfolding adds_def .. hence "t + s = t + (u + v)" unfolding eq1 by (simp add: ac_simps) hence s: "s = u + v" unfolding add_left_cancel . show ?thesis unfolding eq2 gcs_def unfolding s by simp qed lemma gcs_comm: "gcs s t = gcs t s" unfolding gcs_def by (simp add: lcs_comm ac_simps) lemma gcs_adds_2: "(gcs s t) adds t" by (simp only: gcs_comm[of s t], rule gcs_adds) end class ulcs_powerprod = lcs_powerprod + ninv_comm_monoid_add begin lemma adds_antisym: assumes "s adds t" "t adds s" shows "s = t" proof - from \s adds t\ obtain u where u_def: "t = s + u" unfolding adds_def .. from \t adds s\ obtain v where v_def: "s = t + v" unfolding adds_def .. from u_def v_def have "s = (s + u) + v" by (simp add: ac_simps) hence "s + 0 = s + (u + v)" by (simp add: ac_simps) hence "u + v = 0" by simp hence "u = 0" using plus_eq_zero[of u v] by simp thus ?thesis using u_def by simp qed lemma lcs_unique: assumes "s adds l" and "t adds l" and *: "\u. s adds u \ t adds u \ l adds u" shows "l = lcs s t" by (rule adds_antisym, rule *, fact adds_lcs, fact adds_lcs_2, rule lcs_adds, fact+) lemma lcs_zero: "lcs 0 t = t" by (rule lcs_unique[symmetric], fact zero_adds, fact adds_refl) lemma lcs_plus_left: "lcs (u + s) (u + t) = u + lcs s t" proof (rule lcs_unique[symmetric], simp_all only: adds_canc_2, fact adds_lcs, fact adds_lcs_2, simp add: add.commute[of u] plus_adds) fix v assume "u adds v \ s adds v - u" hence "s adds v - u" .. assume "t adds v - u" with \s adds v - u\ show "lcs s t adds v - u" by (rule lcs_adds) qed lemma lcs_plus_right: "lcs (s + u) (t + u) = (lcs s t) + u" using lcs_plus_left[of u s t] by (simp add: ac_simps) lemma adds_gcs: assumes "u adds s" and "u adds t" shows "u adds (gcs s t)" proof - from assms have "s + u adds s + t" and "t + u adds t + s" by (simp_all add: plus_adds_mono) hence "lcs (s + u) (t + u) adds s + t" by (auto intro: lcs_adds simp add: ac_simps) hence "u + (lcs s t) adds s + t" unfolding lcs_plus_right by (simp add: ac_simps) hence "u adds (s + t) - (lcs s t)" unfolding plus_adds .. thus ?thesis unfolding gcs_def . qed lemma gcs_unique: assumes "g adds s" and "g adds t" and *: "\u. u adds s \ u adds t \ u adds g" shows "g = gcs s t" by (rule adds_antisym, rule adds_gcs, fact, fact, rule *, fact gcs_adds, fact gcs_adds_2) lemma gcs_plus_left: "gcs (u + s) (u + t) = u + gcs s t" proof - have "u + s + (u + t) - (u + lcs s t) = u + s + (u + t) - u - lcs s t" by (simp only: diff_diff_add) also have "... = u + s + t + (u - u) - lcs s t" by (simp add: add.left_commute) also have "... = u + s + t - lcs s t" by simp also have "... = u + (s + t - lcs s t)" using add_assoc add_commute local.lcs_adds_plus local.minus_plus by auto finally show ?thesis unfolding gcs_def lcs_plus_left . qed lemma gcs_plus_right: "gcs (s + u) (t + u) = (gcs s t) + u" using gcs_plus_left[of u s t] by (simp add: ac_simps) lemma lcs_same [simp]: "lcs s s = s" proof - have "lcs s s adds s" by (rule lcs_adds, simp_all) moreover have "s adds lcs s s" by (rule adds_lcs) ultimately show ?thesis by (rule adds_antisym) qed lemma gcs_same [simp]: "gcs s s = s" proof - have "gcs s s adds s" by (rule gcs_adds) moreover have "s adds gcs s s" by (rule adds_gcs, simp_all) ultimately show ?thesis by (rule adds_antisym) qed end subsection \Dickson Classes\ definition (in plus) dickson_grading :: "('a \ nat) \ bool" where "dickson_grading d \ ((\s t. d (s + t) = max (d s) (d t)) \ (\n::nat. almost_full_on (adds) {x. d x \ n}))" definition dgrad_set :: "('a \ nat) \ nat \ 'a set" where "dgrad_set d m = {t. d t \ m}" definition dgrad_set_le :: "('a \ nat) \ ('a set) \ ('a set) \ bool" where "dgrad_set_le d S T \ (\s\S. \t\T. d s \ d t)" lemma dickson_gradingI: assumes "\s t. d (s + t) = max (d s) (d t)" assumes "\n::nat. almost_full_on (adds) {x. d x \ n}" shows "dickson_grading d" unfolding dickson_grading_def using assms by blast lemma dickson_gradingD1: "dickson_grading d \ d (s + t) = max (d s) (d t)" by (auto simp add: dickson_grading_def) lemma dickson_gradingD2: "dickson_grading d \ almost_full_on (adds) {x. d x \ n}" by (auto simp add: dickson_grading_def) lemma dickson_gradingD2': assumes "dickson_grading (d::'a::comm_monoid_add \ nat)" shows "wqo_on (adds) {x. d x \ n}" proof (intro wqo_onI transp_onI) fix x y z :: 'a assume "x adds y" and "y adds z" thus "x adds z" by (rule adds_trans) next from assms show "almost_full_on (adds) {x. d x \ n}" by (rule dickson_gradingD2) qed lemma dickson_gradingE: assumes "dickson_grading d" and "\i::nat. d ((seq::nat \ 'a::plus) i) \ n" obtains i j where "i < j" and "seq i adds seq j" proof - from assms(1) have "almost_full_on (adds) {x. d x \ n}" by (rule dickson_gradingD2) moreover from assms(2) have "\i. seq i \ {x. d x \ n}" by simp ultimately obtain i j where "i < j" and "seq i adds seq j" by (rule almost_full_onD) thus ?thesis .. qed lemma dickson_grading_adds_imp_le: assumes "dickson_grading d" and "s adds t" shows "d s \ d t" proof - from assms(2) obtain u where "t = s + u" .. hence "d t = max (d s) (d u)" by (simp only: dickson_gradingD1[OF assms(1)]) thus ?thesis by simp qed lemma dickson_grading_minus: assumes "dickson_grading d" and "s adds (t::'a::cancel_ab_semigroup_add)" shows "d (t - s) \ d t" proof - from assms(2) obtain u where "t = s + u" .. hence "t - s = u" by simp from assms(1) have "d t = ord_class.max (d s) (d u)" unfolding \t = s + u\ by (rule dickson_gradingD1) thus ?thesis by (simp add: \t - s = u\) qed lemma dickson_grading_lcs: assumes "dickson_grading d" shows "d (lcs s t) \ max (d s) (d t)" proof - from assms have "d (lcs s t) \ d (s + t)" by (rule dickson_grading_adds_imp_le, intro lcs_adds_plus) thus ?thesis by (simp only: dickson_gradingD1[OF assms]) qed lemma dickson_grading_lcs_minus: assumes "dickson_grading d" shows "d (lcs s t - s) \ max (d s) (d t)" proof - from assms have "d (lcs s t - s) \ d (lcs s t)" by (rule dickson_grading_minus, intro adds_lcs) also from assms have "... \ max (d s) (d t)" by (rule dickson_grading_lcs) finally show ?thesis . qed lemma dgrad_set_leI: assumes "\s. s \ S \ \t\T. d s \ d t" shows "dgrad_set_le d S T" using assms by (auto simp: dgrad_set_le_def) lemma dgrad_set_leE: assumes "dgrad_set_le d S T" and "s \ S" obtains t where "t \ T" and "d s \ d t" using assms by (auto simp: dgrad_set_le_def) lemma dgrad_set_exhaust_expl: assumes "finite F" shows "F \ dgrad_set d (Max (d ` F))" proof fix f assume "f \ F" hence "d f \ d ` F" by simp with _ have "d f \ Max (d ` F)" proof (rule Max_ge) from assms show "finite (d ` F)" by auto qed hence "dgrad_set d (d f) \ dgrad_set d (Max (d ` F))" by (auto simp: dgrad_set_def) moreover have "f \ dgrad_set d (d f)" by (simp add: dgrad_set_def) ultimately show "f \ dgrad_set d (Max (d ` F))" .. qed lemma dgrad_set_exhaust: assumes "finite F" obtains m where "F \ dgrad_set d m" proof from assms show "F \ dgrad_set d (Max (d ` F))" by (rule dgrad_set_exhaust_expl) qed lemma dgrad_set_le_trans [trans]: assumes "dgrad_set_le d S T" and "dgrad_set_le d T U" shows "dgrad_set_le d S U" unfolding dgrad_set_le_def proof fix s assume "s \ S" with assms(1) obtain t where "t \ T" and 1: "d s \ d t" by (auto simp add: dgrad_set_le_def) from assms(2) this(1) obtain u where "u \ U" and 2: "d t \ d u" by (auto simp add: dgrad_set_le_def) from this(1) show "\u\U. d s \ d u" proof from 1 2 show "d s \ d u" by (rule le_trans) qed qed lemma dgrad_set_le_Un: "dgrad_set_le d (S \ T) U \ (dgrad_set_le d S U \ dgrad_set_le d T U)" by (auto simp add: dgrad_set_le_def) lemma dgrad_set_le_subset: assumes "S \ T" shows "dgrad_set_le d S T" unfolding dgrad_set_le_def using assms by blast lemma dgrad_set_le_refl: "dgrad_set_le d S S" by (rule dgrad_set_le_subset, fact subset_refl) lemma dgrad_set_le_dgrad_set: assumes "dgrad_set_le d F G" and "G \ dgrad_set d m" shows "F \ dgrad_set d m" proof fix f assume "f \ F" with assms(1) obtain g where "g \ G" and *: "d f \ d g" by (auto simp add: dgrad_set_le_def) from assms(2) this(1) have "g \ dgrad_set d m" .. hence "d g \ m" by (simp add: dgrad_set_def) with * have "d f \ m" by (rule le_trans) thus "f \ dgrad_set d m" by (simp add: dgrad_set_def) qed lemma dgrad_set_dgrad: "p \ dgrad_set d (d p)" by (simp add: dgrad_set_def) lemma dgrad_setI [intro]: assumes "d t \ m" shows "t \ dgrad_set d m" using assms by (auto simp: dgrad_set_def) lemma dgrad_setD: assumes "t \ dgrad_set d m" shows "d t \ m" using assms by (simp add: dgrad_set_def) lemma dgrad_set_zero [simp]: "dgrad_set (\_. 0) m = UNIV" by auto lemma subset_dgrad_set_zero: "F \ dgrad_set (\_. 0) m" by simp lemma dgrad_set_subset: assumes "m \ n" shows "dgrad_set d m \ dgrad_set d n" using assms by (auto simp: dgrad_set_def) lemma dgrad_set_closed_plus: assumes "dickson_grading d" and "s \ dgrad_set d m" and "t \ dgrad_set d m" shows "s + t \ dgrad_set d m" proof - from assms(1) have "d (s + t) = ord_class.max (d s) (d t)" by (rule dickson_gradingD1) also from assms(2, 3) have "... \ m" by (simp add: dgrad_set_def) finally show ?thesis by (simp add: dgrad_set_def) qed lemma dgrad_set_closed_minus: assumes "dickson_grading d" and "s \ dgrad_set d m" and "t adds (s::'a::cancel_ab_semigroup_add)" shows "s - t \ dgrad_set d m" proof - from assms(1, 3) have "d (s - t) \ d s" by (rule dickson_grading_minus) also from assms(2) have "... \ m" by (simp add: dgrad_set_def) finally show ?thesis by (simp add: dgrad_set_def) qed lemma dgrad_set_closed_lcs: assumes "dickson_grading d" and "s \ dgrad_set d m" and "t \ dgrad_set d m" shows "lcs s t \ dgrad_set d m" proof - from assms(1) have "d (lcs s t) \ ord_class.max (d s) (d t)" by (rule dickson_grading_lcs) also from assms(2, 3) have "... \ m" by (simp add: dgrad_set_def) finally show ?thesis by (simp add: dgrad_set_def) qed lemma dickson_gradingD_dgrad_set: "dickson_grading d \ almost_full_on (adds) (dgrad_set d m)" by (auto dest: dickson_gradingD2 simp: dgrad_set_def) lemma ex_finite_adds: assumes "dickson_grading d" and "S \ dgrad_set d m" obtains T where "finite T" and "T \ S" and "\s. s \ S \ (\t\T. t adds (s::'a::cancel_comm_monoid_add))" proof - have "reflp ((adds)::'a \ _)" by (simp add: reflp_def) moreover from assms(2) have "almost_full_on (adds) S" proof (rule almost_full_on_subset) from assms(1) show "almost_full_on (adds) (dgrad_set d m)" by (rule dickson_gradingD_dgrad_set) qed ultimately obtain T where "finite T" and "T \ S" and "\s. s \ S \ (\t\T. t adds s)" by (rule almost_full_on_finite_subsetE, blast) thus ?thesis .. qed class graded_dickson_powerprod = ulcs_powerprod + assumes ex_dgrad: "\d::'a \ nat. dickson_grading d" begin definition dgrad_dummy where "dgrad_dummy = (SOME d. dickson_grading d)" lemma dickson_grading_dgrad_dummy: "dickson_grading dgrad_dummy" unfolding dgrad_dummy_def using ex_dgrad by (rule someI_ex) end (* graded_dickson_powerprod *) class dickson_powerprod = ulcs_powerprod + assumes dickson: "almost_full_on (adds) UNIV" begin lemma dickson_grading_zero: "dickson_grading (\_::'a. 0)" by (simp add: dickson_grading_def dickson) subclass graded_dickson_powerprod by (standard, rule, fact dickson_grading_zero) end (* dickson_powerprod *) text \Class @{class graded_dickson_powerprod} is a slightly artificial construction. It is needed, because type @{typ "nat \\<^sub>0 nat"} does not satisfy the usual conditions of a "Dickson domain" (as formulated in class @{class dickson_powerprod}), but we still want to use that type as the type of power-products in the computation of Gr\"obner bases. So, we exploit the fact that in a finite set of polynomials (which is the input of Buchberger's algorithm) there is always some "highest" indeterminate that occurs with non-zero exponent, and no "higher" indeterminates are generated during the execution of the algorithm. This allows us to prove that the algorithm terminates, even though there are in principle infinitely many indeterminates.\ subsection \Additive Linear Orderings\ lemma group_eq_aux: "a + (b - a) = (b::'a::ab_group_add)" proof - have "a + (b - a) = b - a + a" by simp also have "... = b" by simp finally show ?thesis . qed class semi_canonically_ordered_monoid_add = ordered_comm_monoid_add + assumes le_imp_add: "a \ b \ (\c. b = a + c)" context canonically_ordered_monoid_add begin subclass semi_canonically_ordered_monoid_add by (standard, simp only: le_iff_add) end class add_linorder_group = ordered_ab_semigroup_add_imp_le + ab_group_add + linorder class add_linorder = ordered_ab_semigroup_add_imp_le + cancel_comm_monoid_add + semi_canonically_ordered_monoid_add + linorder begin subclass ordered_comm_monoid_add .. subclass ordered_cancel_comm_monoid_add .. lemma le_imp_inv: assumes "a \ b" shows "b = a + (b - a)" using le_imp_add[OF assms] by auto lemma max_eq_sum: obtains y where "max a b = a + y" unfolding max_def proof (cases "a \ b") case True hence "b = a + (b - a)" by (rule le_imp_inv) then obtain c where eq: "b = a + c" .. show ?thesis proof from True show "max a b = a + c" unfolding max_def eq by simp qed next case False show ?thesis proof from False show "max a b = a + 0" unfolding max_def by simp qed qed lemma min_plus_max: shows "(min a b) + (max a b) = a + b" proof (cases "a \ b") case True thus ?thesis unfolding min_def max_def by simp next case False thus ?thesis unfolding min_def max_def by (simp add: ac_simps) qed end (* add_linorder *) class add_linorder_min = add_linorder + assumes zero_min: "0 \ x" begin subclass ninv_comm_monoid_add proof fix x y assume *: "x + y = 0" show "x = 0" proof - from zero_min[of x] have "0 = x \ x > 0" by auto thus ?thesis proof assume "x > 0" have "0 \ y" by (fact zero_min) also have "... = 0 + y" by simp also from \x > 0\ have "... < x + y" by (rule add_strict_right_mono) finally have "0 < x + y" . hence "x + y \ 0" by simp from this * show ?thesis .. qed simp qed qed lemma leq_add_right: shows "x \ x + y" using add_left_mono[OF zero_min[of y], of x] by simp lemma leq_add_left: shows "x \ y + x" using add_right_mono[OF zero_min[of y], of x] by simp subclass canonically_ordered_monoid_add by (standard, rule, elim le_imp_add, elim exE, simp add: leq_add_right) end (* add_linorder_min *) class add_wellorder = add_linorder_min + wellorder instantiation nat :: add_linorder begin instance by (standard, simp) end (* add_linorder *) instantiation nat :: add_linorder_min begin instance by (standard, simp) end instantiation nat :: add_wellorder begin instance .. end context add_linorder_group begin subclass add_linorder proof (standard, intro exI) fix a b show "b = a + (b - a)" using add_commute local.diff_add_cancel by auto qed end (* add_linorder_group *) instantiation int :: add_linorder_group begin instance .. end instantiation rat :: add_linorder_group begin instance .. end instantiation real :: add_linorder_group begin instance .. end subsection \Ordered Power-Products\ locale ordered_powerprod = ordered_powerprod_lin: linorder ord ord_strict for ord::"'a \ 'a::comm_powerprod \ bool" (infixl "\" 50) and ord_strict::"'a \ 'a::comm_powerprod \ bool" (infixl "\" 50) + assumes zero_min: "0 \ t" assumes plus_monotone: "s \ t \ s + u \ t + u" begin abbreviation ord_conv (infixl "\" 50) where "ord_conv \ (\)\\" abbreviation ord_strict_conv (infixl "\" 50) where "ord_strict_conv \ (\)\\" lemma ord_canc: assumes "s + u \ t + u" shows "s \ t" proof (rule ordered_powerprod_lin.le_cases[of s t], simp) assume "t \ s" from assms plus_monotone[OF this, of u] have "t + u = s + u" using ordered_powerprod_lin.order.eq_iff by simp hence "t = s" by simp thus "s \ t" by simp qed lemma ord_adds: assumes "s adds t" shows "s \ t" proof - from assms have "\u. t = s + u" unfolding adds_def by simp then obtain k where "t = s + k" .. thus ?thesis using plus_monotone[OF zero_min[of k], of s] by (simp add: ac_simps) qed lemma ord_canc_left: assumes "u + s \ u + t" shows "s \ t" using assms unfolding add.commute[of u] by (rule ord_canc) lemma ord_strict_canc: assumes "s + u \ t + u" shows "s \ t" using assms ord_canc[of s u t] add_right_cancel[of s u t] ordered_powerprod_lin.le_imp_less_or_eq ordered_powerprod_lin.order.strict_implies_order by blast lemma ord_strict_canc_left: assumes "u + s \ u + t" shows "s \ t" using assms unfolding add.commute[of u] by (rule ord_strict_canc) lemma plus_monotone_left: assumes "s \ t" shows "u + s \ u + t" using assms by (simp add: add.commute, rule plus_monotone) lemma plus_monotone_strict: assumes "s \ t" shows "s + u \ t + u" using assms by (simp add: ordered_powerprod_lin.order.strict_iff_order plus_monotone) lemma plus_monotone_strict_left: assumes "s \ t" shows "u + s \ u + t" using assms by (simp add: ordered_powerprod_lin.order.strict_iff_order plus_monotone_left) end locale gd_powerprod = ordered_powerprod ord ord_strict for ord::"'a \ 'a::graded_dickson_powerprod \ bool" (infixl "\" 50) and ord_strict (infixl "\" 50) begin definition dickson_le :: "('a \ nat) \ nat \ 'a \ 'a \ bool" where "dickson_le d m s t \ (d s \ m \ d t \ m \ s \ t)" definition dickson_less :: "('a \ nat) \ nat \ 'a \ 'a \ bool" where "dickson_less d m s t \ (d s \ m \ d t \ m \ s \ t)" lemma dickson_leI: assumes "d s \ m" and "d t \ m" and "s \ t" shows "dickson_le d m s t" using assms by (simp add: dickson_le_def) lemma dickson_leD1: assumes "dickson_le d m s t" shows "d s \ m" using assms by (simp add: dickson_le_def) lemma dickson_leD2: assumes "dickson_le d m s t" shows "d t \ m" using assms by (simp add: dickson_le_def) lemma dickson_leD3: assumes "dickson_le d m s t" shows "s \ t" using assms by (simp add: dickson_le_def) lemma dickson_le_trans: assumes "dickson_le d m s t" and "dickson_le d m t u" shows "dickson_le d m s u" using assms by (auto simp add: dickson_le_def) lemma dickson_lessI: assumes "d s \ m" and "d t \ m" and "s \ t" shows "dickson_less d m s t" using assms by (simp add: dickson_less_def) lemma dickson_lessD1: assumes "dickson_less d m s t" shows "d s \ m" using assms by (simp add: dickson_less_def) lemma dickson_lessD2: assumes "dickson_less d m s t" shows "d t \ m" using assms by (simp add: dickson_less_def) lemma dickson_lessD3: assumes "dickson_less d m s t" shows "s \ t" using assms by (simp add: dickson_less_def) lemma dickson_less_irrefl: "\ dickson_less d m t t" by (simp add: dickson_less_def) lemma dickson_less_trans: assumes "dickson_less d m s t" and "dickson_less d m t u" shows "dickson_less d m s u" using assms by (auto simp add: dickson_less_def) lemma transp_dickson_less: "transp (dickson_less d m)" by (rule transpI, fact dickson_less_trans) lemma wfp_on_ord_strict: assumes "dickson_grading d" shows "wfp_on (\) {x. d x \ n}" proof - let ?A = "{x. d x \ n}" have "strict (\) = (\)" by (intro ext, simp only: ordered_powerprod_lin.less_le_not_le) have "qo_on (adds) ?A" by (auto simp: qo_on_def reflp_on_def transp_on_def dest: adds_trans) moreover from assms have "wqo_on (adds) ?A" by (rule dickson_gradingD2') ultimately have "(\Q. (\x\?A. \y\?A. x adds y \ Q x y) \ qo_on Q ?A \ wfp_on (strict Q) ?A)" by (simp only: wqo_extensions_wf_conv) hence "(\x\?A. \y\?A. x adds y \ x \ y) \ qo_on (\) ?A \ wfp_on (strict (\)) ?A" .. thus ?thesis unfolding \strict (\) = (\)\ proof show "(\x\?A. \y\?A. x adds y \ x \ y) \ qo_on (\) ?A" proof (intro conjI ballI impI ord_adds) show "qo_on (\) ?A" by (auto simp: qo_on_def reflp_on_def transp_on_def) qed qed qed lemma wf_dickson_less: assumes "dickson_grading d" shows "wfP (dickson_less d m)" proof (rule wfP_chain) show "\ (\seq. \i. dickson_less d m (seq (Suc i)) (seq i))" proof assume "\seq. \i. dickson_less d m (seq (Suc i)) (seq i)" then obtain seq::"nat \ 'a" where "\i. dickson_less d m (seq (Suc i)) (seq i)" .. hence *: "\i. dickson_less d m (seq (Suc i)) (seq i)" .. with transp_dickson_less have seq_decr: "\i j. i < j \ dickson_less d m (seq j) (seq i)" by (rule transp_sequence) from assms obtain i j where "i < j" and i_adds_j: "seq i adds seq j" proof (rule dickson_gradingE) fix i from * show "d (seq i) \ m" by (rule dickson_lessD2) qed from \i < j\ have "dickson_less d m (seq j) (seq i)" by (rule seq_decr) hence "seq j \ seq i" by (rule dickson_lessD3) moreover from i_adds_j have "seq i \ seq j" by (rule ord_adds) ultimately show False by simp qed qed end text \\gd_powerprod\ stands for @{emph \graded ordered Dickson power-products\}.\ locale od_powerprod = ordered_powerprod ord ord_strict for ord::"'a \ 'a::dickson_powerprod \ bool" (infixl "\" 50) and ord_strict (infixl "\" 50) begin sublocale gd_powerprod by standard lemma wf_ord_strict: "wfP (\)" proof (rule wfP_chain) show "\ (\seq. \i. seq (Suc i) \ seq i)" proof assume "\seq. \i. seq (Suc i) \ seq i" then obtain seq::"nat \ 'a" where "\i. seq (Suc i) \ seq i" .. hence "\i. seq (Suc i) \ seq i" .. - with ordered_powerprod_lin.transp_less have seq_decr: "\i j. i < j \ (seq j) \ (seq i)" + with ordered_powerprod_lin.transp_on_less have seq_decr: "\i j. i < j \ (seq j) \ (seq i)" by (rule transp_sequence) from dickson obtain i j::nat where "i < j" and i_adds_j: "seq i adds seq j" by (auto elim!: almost_full_onD) from seq_decr[OF \i < j\] have "seq j \ seq i \ seq j \ seq i" by auto hence "seq j \ seq i" and "seq j \ seq i" by simp_all from \seq j \ seq i\ \seq j \ seq i\ ord_adds[OF i_adds_j] ordered_powerprod_lin.order.eq_iff[of "seq j" "seq i"] show False by simp qed qed end text \\od_powerprod\ stands for @{emph \ordered Dickson power-products\}.\ subsection \Functions as Power-Products\ lemma finite_neq_0: assumes fin_A: "finite {x. f x \ 0}" and fin_B: "finite {x. g x \ 0}" and "\x. h x 0 0 = 0" shows "finite {x. h x (f x) (g x) \ 0}" proof - from fin_A fin_B have "finite ({x. f x \ 0} \ {x. g x \ 0})" by (intro finite_UnI) hence finite_union: "finite {x. (f x \ 0) \ (g x \ 0)}" by (simp only: Collect_disj_eq) have "{x. h x (f x) (g x) \ 0} \ {x. (f x \ 0) \ (g x \ 0)}" proof (intro Collect_mono, rule) fix x::'a assume h_not_zero: "h x (f x) (g x) \ 0" have "f x = 0 \ g x \ 0" proof assume "f x = 0" "g x = 0" thus False using h_not_zero \h x 0 0 = 0\ by simp qed thus "f x \ 0 \ g x \ 0" by auto qed from finite_subset[OF this] finite_union show "finite {x. h x (f x) (g x) \ 0}" . qed lemma finite_neq_0': assumes "finite {x. f x \ 0}" and "finite {x. g x \ 0}" and "h 0 0 = 0" shows "finite {x. h (f x) (g x) \ 0}" using assms by (rule finite_neq_0) lemma finite_neq_0_inv: assumes fin_A: "finite {x. h x (f x) (g x) \ 0}" and fin_B: "finite {x. f x \ 0}" and "\x y. h x 0 y = y" shows "finite {x. g x \ 0}" proof - from fin_A and fin_B have "finite ({x. h x (f x) (g x) \ 0} \ {x. f x \ 0})" by (intro finite_UnI) hence finite_union: "finite {x. (h x (f x) (g x) \ 0) \ f x \ 0}" by (simp only: Collect_disj_eq) have "{x. g x \ 0} \ {x. (h x (f x) (g x) \ 0) \ f x \ 0}" by (intro Collect_mono, rule, rule disjCI, simp add: assms(3)) from this finite_union show "finite {x. g x \ 0}" by (rule finite_subset) qed lemma finite_neq_0_inv': assumes inf_A: "finite {x. h (f x) (g x) \ 0}" and fin_B: "finite {x. f x \ 0}" and "\x. h 0 x = x" shows "finite {x. g x \ 0}" using assms by (rule finite_neq_0_inv) subsubsection \@{typ "'a \ 'b"} belongs to class @{class comm_powerprod}\ instance "fun" :: (type, cancel_comm_monoid_add) comm_powerprod by standard subsubsection \@{typ "'a \ 'b"} belongs to class @{class ninv_comm_monoid_add}\ instance "fun" :: (type, ninv_comm_monoid_add) ninv_comm_monoid_add by (standard, simp only: plus_fun_def zero_fun_def fun_eq_iff, intro allI, rule plus_eq_zero, auto) subsubsection \@{typ "'a \ 'b"} belongs to class @{class lcs_powerprod}\ instantiation "fun" :: (type, add_linorder) lcs_powerprod begin definition lcs_fun::"('a \ 'b) \ ('a \ 'b) \ ('a \ 'b)" where "lcs f g = (\x. max (f x) (g x))" lemma adds_funI: assumes "s \ t" shows "s adds (t::'a \ 'b)" proof (rule addsI, rule) fix x from assms have "s x \ t x" unfolding le_fun_def .. hence "t x = s x + (t x - s x)" by (rule le_imp_inv) thus "t x = (s + (t - s)) x" by simp qed lemma adds_fun_iff: "f adds (g::'a \ 'b) \ (\x. f x adds g x)" unfolding adds_def plus_fun_def by metis lemma adds_fun_iff': "f adds (g::'a \ 'b) \ (\x. \y. g x = f x + y)" unfolding adds_fun_iff unfolding adds_def plus_fun_def .. lemma adds_lcs_fun: shows "s adds (lcs s (t::'a \ 'b))" by (rule adds_funI, simp only: le_fun_def lcs_fun_def, auto simp: max_def) lemma lcs_comm_fun: "lcs s t = lcs t (s::'a \ 'b)" unfolding lcs_fun_def by (auto simp: max_def intro!: ext) lemma lcs_adds_fun: assumes "s adds u" and "t adds (u::'a \ 'b)" shows "(lcs s t) adds u" using assms unfolding lcs_fun_def adds_fun_iff' proof - assume a1: "\x. \y. u x = s x + y" and a2: "\x. \y. u x = t x + y" show "\x. \y. u x = max (s x) (t x) + y" proof fix x from a1 have b1: "\y. u x = s x + y" .. from a2 have b2: "\y. u x = t x + y" .. show "\y. u x = max (s x) (t x) + y" unfolding max_def by (split if_split, intro conjI impI, rule b2, rule b1) qed qed instance apply standard subgoal by (rule adds_lcs_fun) subgoal by (rule lcs_adds_fun) subgoal by (rule lcs_comm_fun) done end lemma leq_lcs_fun_1: "s \ (lcs s (t::'a \ 'b::add_linorder))" by (simp add: lcs_fun_def le_fun_def) lemma leq_lcs_fun_2: "t \ (lcs s (t::'a \ 'b::add_linorder))" by (simp add: lcs_fun_def le_fun_def) lemma lcs_leq_fun: assumes "s \ u" and "t \ (u::'a \ 'b::add_linorder)" shows "(lcs s t) \ u" using assms by (simp add: lcs_fun_def le_fun_def) lemma adds_fun: "s adds t \ s \ t" for s t::"'a \ 'b::add_linorder_min" proof assume "s adds t" from this obtain k where "t = s + k" .. show "s \ t" unfolding \t = s + k\ le_fun_def plus_fun_def le_iff_add by (simp add: leq_add_right) qed (rule adds_funI) lemma gcs_fun: "gcs s (t::'a \ ('b::add_linorder)) = (\x. min (s x) (t x))" proof - show ?thesis unfolding gcs_def lcs_fun_def fun_diff_def proof (simp, rule) fix x have eq: "s x + t x = max (s x) (t x) + min (s x) (t x)" by (metis add.commute min_def max_def) thus "s x + t x - max (s x) (t x) = min (s x) (t x)" by simp qed qed lemma gcs_leq_fun_1: "(gcs s (t::'a \ 'b::add_linorder)) \ s" by (simp add: gcs_fun le_fun_def) lemma gcs_leq_fun_2: "(gcs s (t::'a \ 'b::add_linorder)) \ t" by (simp add: gcs_fun le_fun_def) lemma leq_gcs_fun: assumes "u \ s" and "u \ (t::'a \ 'b::add_linorder)" shows "u \ (gcs s t)" using assms by (simp add: gcs_fun le_fun_def) subsubsection \@{typ "'a \ 'b"} belongs to class @{class ulcs_powerprod}\ instance "fun" :: (type, add_linorder_min) ulcs_powerprod .. subsubsection \Power-products in a given set of indeterminates\ definition supp_fun::"('a \ 'b::zero) \ 'a set" where "supp_fun f = {x. f x \ 0}" text \@{term supp_fun} for general functions is like @{term keys} for @{type poly_mapping}, but does not need to be finite.\ lemma keys_eq_supp: "keys s = supp_fun (lookup s)" unfolding supp_fun_def by (transfer, rule) lemma supp_fun_zero [simp]: "supp_fun 0 = {}" by (auto simp: supp_fun_def) lemma supp_fun_eq_zero_iff: "supp_fun f = {} \ f = 0" by (auto simp: supp_fun_def) lemma sub_supp_empty: "supp_fun s \ {} \ (s = 0)" by (auto simp: supp_fun_def) lemma except_fun_idI: "supp_fun f \ V = {} \ except_fun f V = f" by (auto simp: except_fun_def supp_fun_def when_def intro!: ext) lemma supp_except_fun: "supp_fun (except_fun s V) = supp_fun s - V" by (auto simp: except_fun_def supp_fun_def) lemma supp_fun_plus_subset: "supp_fun (s + t) \ supp_fun s \ supp_fun (t::'a \ 'b::monoid_add)" unfolding supp_fun_def by force lemma fun_eq_zeroI: assumes "\x. x \ supp_fun f \ f x = 0" shows "f = 0" proof (rule, simp) fix x show "f x = 0" proof (cases "x \ supp_fun f") case True then show ?thesis by (rule assms) next case False then show ?thesis by (simp add: supp_fun_def) qed qed lemma except_fun_cong1: "supp_fun s \ ((V - U) \ (U - V)) \ {} \ except_fun s V = except_fun s U" by (auto simp: except_fun_def when_def supp_fun_def intro!: ext) lemma adds_except_fun: "s adds t = (except_fun s V adds except_fun t V \ except_fun s (- V) adds except_fun t (- V))" for s t :: "'a \ 'b::add_linorder" by (auto simp: supp_fun_def except_fun_def adds_fun_iff when_def) lemma adds_except_fun_singleton: "s adds t = (except_fun s {v} adds except_fun t {v} \ s v adds t v)" for s t :: "'a \ 'b::add_linorder" by (auto simp: supp_fun_def except_fun_def adds_fun_iff when_def) subsubsection \Dickson's lemma for power-products in finitely many indeterminates\ lemma Dickson_fun: assumes "finite V" shows "almost_full_on (adds) {x::'a \ 'b::add_wellorder. supp_fun x \ V}" using assms proof (induct V) case empty have "finite {0}" by simp moreover have "reflp_on {0::'a \ 'b} (adds)" by (simp add: reflp_on_def) ultimately have "almost_full_on (adds) {0::'a \ 'b}" by (rule finite_almost_full_on) thus ?case by (simp add: supp_fun_eq_zero_iff) next case (insert v V) show ?case proof (rule almost_full_onI) fix seq::"nat \ 'a \ 'b" assume "\i. seq i \ {x. supp_fun x \ insert v V}" hence a: "supp_fun (seq i) \ insert v V" for i by simp define seq' where "seq' = (\i. (except_fun (seq i) {v}, except_fun (seq i) V))" have "almost_full_on (adds) {x::'a \ 'b. supp_fun x \ {v}}" proof (rule almost_full_onI) fix f::"nat \ 'a \ 'b" assume "\i. f i \ {x. supp_fun x \ {v}}" hence b: "supp_fun (f i) \ {v}" for i by simp let ?f = "\i. f i v" have "wfP ((<)::'b \ _)" by (simp add: wf wfP_def) hence "\f::nat \ 'b. \i. f i \ f (Suc i)" by (simp add: wf_iff_no_infinite_down_chain[to_pred] not_less) hence "\i. ?f i \ ?f (Suc i)" .. then obtain i where "?f i \ ?f (Suc i)" .. have "i < Suc i" by simp moreover have "f i adds f (Suc i)" unfolding adds_fun_iff proof fix x show "f i x adds f (Suc i) x" proof (cases "x = v") case True with \?f i \ ?f (Suc i)\ show ?thesis by (simp add: adds_def le_iff_add) next case False with b have "x \ supp_fun (f i)" and "x \ supp_fun (f (Suc i))" by blast+ thus ?thesis by (simp add: supp_fun_def) qed qed ultimately show "good (adds) f" by (meson goodI) qed with insert(3) have "almost_full_on (prod_le (adds) (adds)) ({x::'a \ 'b. supp_fun x \ V} \ {x::'a \ 'b. supp_fun x \ {v}})" (is "almost_full_on ?P ?A") by (rule almost_full_on_Sigma) moreover from a have "seq' i \ ?A" for i by (auto simp add: seq'_def supp_except_fun) ultimately obtain i j where "i < j" and "?P (seq' i) (seq' j)" by (rule almost_full_onD) have "seq i adds seq j" unfolding adds_except_fun[where s="seq i" and V=V] proof from \?P (seq' i) (seq' j)\ show "except_fun (seq i) V adds except_fun (seq j) V" by (simp add: prod_le_def seq'_def) next from \?P (seq' i) (seq' j)\ have "except_fun (seq i) {v} adds except_fun (seq j) {v}" by (simp add: prod_le_def seq'_def) moreover have "except_fun (seq i) (- V) = except_fun (seq i) {v}" by (rule except_fun_cong1; use a[of i] insert.hyps(2) in blast) moreover have "except_fun (seq j) (- V) = except_fun (seq j) {v}" by (rule except_fun_cong1; use a[of j] insert.hyps(2) in blast) ultimately show "except_fun (seq i) (- V) adds except_fun (seq j) (- V)" by simp qed with \i < j\ show "good (adds) seq" by (meson goodI) qed qed instance "fun" :: (finite, add_wellorder) dickson_powerprod proof have "finite (UNIV::'a set)" by simp hence "almost_full_on (adds) {x::'a \ 'b. supp_fun x \ UNIV}" by (rule Dickson_fun) thus "almost_full_on (adds) (UNIV::('a \ 'b) set)" by simp qed subsubsection \Lexicographic Term Order\ text \Term orders are certain linear orders on power-products, satisfying additional requirements. Further information on term orders can be found, e.\,g., in @{cite Robbiano1985}.\ context wellorder begin lemma neq_fun_alt: assumes "s \ (t::'a \ 'b)" obtains x where "s x \ t x" and "\y. s y \ t y \ x \ y" proof - from assms ext[of s t] have "\x. s x \ t x" by auto with exists_least_iff[of "\x. s x \ t x"] obtain x where x1: "s x \ t x" and x2: "\y. y < x \ s y = t y" by auto show ?thesis proof from x1 show "s x \ t x" . next fix y assume "s y \ t y" with x2[of y] have "\ y < x" by auto thus "x \ y" by simp qed qed definition lex_fun::"('a \ 'b) \ ('a \ 'b::order) \ bool" where "lex_fun s t \ (\x. s x \ t x \ (\y t y))" definition "lex_fun_strict s t \ lex_fun s t \ \ lex_fun t s" text \Attention! @{term lex_fun} reverses the order of the indeterminates: if @{term x} is smaller than @{term y} w.r.t. the order on @{typ 'a}, then the @{emph \power-product\} @{term x} is @{emph \greater\} than the @{emph \power-product\} @{term y}.\ lemma lex_fun_alt: shows "lex_fun s t = (s = t \ (\x. s x < t x \ (\y t" from neq_fun_alt[OF this] obtain x0 where x0_neq: "s x0 \ t x0" and x0_min: "\z. s z \ t z \ x0 \ z" by auto show ?R proof (intro disjI2, rule exI[of _ x0], intro conjI) from \?L\ have "s x0 \ t x0 \ (\y. y < x0 \ s y \ t y)" unfolding lex_fun_def .. thus "s x0 < t x0" proof assume "s x0 \ t x0" from this x0_neq show ?thesis by simp next assume "\y. y < x0 \ s y \ t y" then obtain y where "y < x0" and y_neq: "s y \ t y" by auto from \y < x0\ x0_min[OF y_neq] show ?thesis by simp qed next show "\y x0 \ y" by simp from this x0_min[of y] show "s y = t y" by auto qed qed qed next assume ?R thus ?L proof assume "s = t" thus ?thesis by (simp add: lex_fun_def) next assume "\x. s x < t x \ (\yz t x \ (\y t y)" proof (cases "s x \ t x") assume "s x \ t x" thus ?thesis by simp next assume x: "\ s x \ t x" show ?thesis proof (intro disjI2, rule exI[of _ y], intro conjI) have "\ x \ y" proof assume "x \ y" hence "x < y \ y = x" by auto thus False proof assume "x < y" from x y_min[rule_format, OF this] show ?thesis by simp next assume "y = x" from this x y show ?thesis by (auto simp: preorder_class.less_le_not_le) qed qed thus "y < x" by simp next from y show "s y \ t y" by simp qed qed qed qed qed lemma lex_fun_refl: "lex_fun s s" unfolding lex_fun_alt by simp lemma lex_fun_antisym: assumes "lex_fun s t" and "lex_fun t s" shows "s = t" proof fix x from assms(1) have "s = t \ (\x. s x < t x \ (\yx. s x < t x \ (\yy (\x. t x < s x \ (\yx. t x < s x \ (\yy x1 < x0 \ x1 = x0" using local.antisym_conv3 by auto show ?thesis proof (rule linorder_cases[of x0 x1]) assume "x1 < x0" from x0_min[rule_format, OF this] x1 show ?thesis by simp next assume "x0 = x1" from this x0 x1 show ?thesis by simp next assume "x0 < x1" from x1_min[rule_format, OF this] x0 show ?thesis by simp qed qed qed qed lemma lex_fun_trans: assumes "lex_fun s t" and "lex_fun t u" shows "lex_fun s u" proof - from assms(1) have "s = t \ (\x. s x < t x \ (\yx. s x < t x \ (\yy (\x. t x < u x \ (\yx. t x < u x \ (\yyx. s x < u x \ (\yx1 < x0\] x1 show "s x1 < u x1" by simp next show "\yx1 < x0\ have "y < x0" by simp from x0_min[rule_format, OF this] x1_min[rule_format, OF \y < x1\] show "s y = u y" by simp qed qed next assume "x0 < x1" show ?thesis proof (rule exI[of _ x0], intro conjI) from x1_min[rule_format, OF \x0 < x1\] x0 show "s x0 < u x0" by simp next show "\yx0 < x1\ have "y < x1" by simp from x0_min[rule_format, OF \y < x0\] x1_min[rule_format, OF this] show "s y = u y" by simp qed qed next assume "x0 = x1" show ?thesis proof (rule exI[of _ x1], intro conjI) from \x0 = x1\ x0 x1 show "s x1 < u x1" by simp next show "\yx0 = x1\ by simp from x0_min[rule_format, OF this] x1_min[rule_format, OF \y < x1\] show "s y = u y" by simp qed qed qed qed qed qed qed lemma lex_fun_lin: "lex_fun s t \ lex_fun t s" for s t::"'a \ 'b::{ordered_comm_monoid_add, linorder}" proof (intro disjCI) assume "\ lex_fun t s" hence a: "\x. \ (t x < s x) \ (\y s y)" unfolding lex_fun_alt by auto show "lex_fun s t" unfolding lex_fun_def proof fix x from a have "\ (t x < s x) \ (\y s y)" .. thus "s x \ t x \ (\y t y)" by auto qed qed corollary lex_fun_strict_alt [code]: "lex_fun_strict s t = (\ lex_fun t s)" for s t::"'a \ 'b::{ordered_comm_monoid_add, linorder}" unfolding lex_fun_strict_def using lex_fun_lin[of s t] by auto lemma lex_fun_zero_min: "lex_fun 0 s" for s::"'a \ 'b::add_linorder_min" by (simp add: lex_fun_def zero_min) lemma lex_fun_plus_monotone: "lex_fun (s + u) (t + u)" if "lex_fun s t" for s t::"'a \ 'b::ordered_cancel_comm_monoid_add" unfolding lex_fun_def proof fix x from that have "s x \ t x \ (\y t y)" unfolding lex_fun_def .. thus "(s + u) x \ (t + u) x \ (\y (t + u) y)" proof assume a1: "s x \ t x" show ?thesis proof (intro disjI1) from a1 show "(s + u) x \ (t + u) x" by (auto simp: add_right_mono) qed next assume "\y t y" then obtain y where "y < x" and a2: "s y \ t y" by auto show ?thesis proof (intro disjI2, rule exI[of _ y], intro conjI, fact) from a2 show "(s + u) y \ (t + u) y" by (auto simp: add_right_mono) qed qed qed end (* wellorder *) subsubsection \Degree\ definition deg_fun::"('a \ 'b::comm_monoid_add) \ 'b" where "deg_fun s \ \x\(supp_fun s). s x" lemma deg_fun_zero[simp]: "deg_fun 0 = 0" by (auto simp: deg_fun_def) lemma deg_fun_eq_0_iff: assumes "finite (supp_fun (s::'a \ 'b::add_linorder_min))" shows "deg_fun s = 0 \ s = 0" proof assume "deg_fun s = 0" hence *: "(\x\(supp_fun s). s x) = 0" by (simp only: deg_fun_def) have **: "\x. 0 \ s x" by (rule zero_min) from * have "\x. x \ supp_fun s \ s x = 0" by (simp only: sum_nonneg_eq_0_iff[OF assms **]) thus "s = 0" by (rule fun_eq_zeroI) qed simp lemma deg_fun_superset: fixes A::"'a set" assumes "supp_fun s \ A" and "finite A" shows "deg_fun s = (\x\A. s x)" unfolding deg_fun_def proof (rule sum.mono_neutral_cong_left, fact, fact, rule) fix x assume "x \ A - supp_fun s" hence "x \ supp_fun s" by simp thus "s x = 0" by (simp add: supp_fun_def) qed rule lemma deg_fun_plus: assumes "finite (supp_fun s)" and "finite (supp_fun t)" shows "deg_fun (s + t) = deg_fun s + deg_fun (t::'a \ 'b::comm_monoid_add)" proof - from assms have fin: "finite (supp_fun s \ supp_fun t)" by simp have "deg_fun (s + t) = (\x\(supp_fun (s + t)). s x + t x)" by (simp add: deg_fun_def) also from fin have "... = (\x\(supp_fun s \ supp_fun t). s x + t x)" proof (rule sum.mono_neutral_cong_left) show "\x\supp_fun s \ supp_fun t - supp_fun (s + t). s x + t x = 0" proof fix x assume "x \ supp_fun s \ supp_fun t - supp_fun (s + t)" hence "x \ supp_fun (s + t)" by simp thus "s x + t x = 0" by (simp add: supp_fun_def) qed qed (rule supp_fun_plus_subset, rule) also have "\ = (\x\(supp_fun s \ supp_fun t). s x) + (\x\(supp_fun s \ supp_fun t). t x)" by (rule sum.distrib) also from fin have "(\x\(supp_fun s \ supp_fun t). s x) = deg_fun s" unfolding deg_fun_def proof (rule sum.mono_neutral_cong_right) show "\x\supp_fun s \ supp_fun t - supp_fun s. s x = 0" proof fix x assume "x \ supp_fun s \ supp_fun t - supp_fun s" hence "x \ supp_fun s" by simp thus "s x = 0" by (simp add: supp_fun_def) qed qed simp_all also from fin have "(\x\(supp_fun s \ supp_fun t). t x) = deg_fun t" unfolding deg_fun_def proof (rule sum.mono_neutral_cong_right) show "\x\supp_fun s \ supp_fun t - supp_fun t. t x = 0" proof fix x assume "x \ supp_fun s \ supp_fun t - supp_fun t" hence "x \ supp_fun t" by simp thus "t x = 0" by (simp add: supp_fun_def) qed qed simp_all finally show ?thesis . qed lemma deg_fun_leq: assumes "finite (supp_fun s)" and "finite (supp_fun t)" and "s \ (t::'a \ 'b::ordered_comm_monoid_add)" shows "deg_fun s \ deg_fun t" proof - let ?A = "supp_fun s \ supp_fun t" from assms(1) assms(2) have 1: "finite ?A" by simp have s: "supp_fun s \ ?A" and t: "supp_fun t \ ?A" by simp_all show ?thesis unfolding deg_fun_superset[OF s 1] deg_fun_superset[OF t 1] proof (rule sum_mono) fix i from assms(3) show "s i \ t i" unfolding le_fun_def .. qed qed subsubsection \General Degree-Orders\ context linorder begin lemma ex_min: assumes "finite (A::'a set)" and "A \ {}" shows "\y\A. (\z\A. y \ z)" using assms proof (induct rule: finite_induct) assume "{} \ {}" thus "\y\{}. \z\{}. y \ z" by simp next fix a::'a and A::"'a set" assume "a \ A" and IH: "A \ {} \ \y\A. (\z\A. y \ z)" show "\y\insert a A. (\z\insert a A. y \ z)" proof (cases "A = {}") case True show ?thesis proof (rule bexI[of _ a], intro ballI) fix z assume "z \ insert a A" from this True have "z = a" by simp thus "a \ z" by simp qed (simp) next case False from IH[OF False] obtain y where "y \ A" and y_min: "\z\A. y \ z" by auto from linear[of a y] show ?thesis proof assume "y \ a" show ?thesis proof (rule bexI[of _ y], intro ballI) fix z assume "z \ insert a A" hence "z = a \ z \ A" by simp thus "y \ z" proof assume "z = a" from this \y \ a\ show "y \ z" by simp next assume "z \ A" from y_min[rule_format, OF this] show "y \ z" . qed next from \y \ A\ show "y \ insert a A" by simp qed next assume "a \ y" show ?thesis proof (rule bexI[of _ a], intro ballI) fix z assume "z \ insert a A" hence "z = a \ z \ A" by simp thus "a \ z" proof assume "z = a" from this show "a \ z" by simp next assume "z \ A" from y_min[rule_format, OF this] \a \ y\ show "a \ z" by simp qed qed (simp) qed qed qed definition dord_fun::"(('a \ 'b::ordered_comm_monoid_add) \ ('a \ 'b) \ bool) \ ('a \ 'b) \ ('a \ 'b) \ bool" where "dord_fun ord s t \ (let d1 = deg_fun s; d2 = deg_fun t in (d1 < d2 \ (d1 = d2 \ ord s t)))" lemma dord_fun_degD: assumes "dord_fun ord s t" shows "deg_fun s \ deg_fun t" using assms unfolding dord_fun_def Let_def by auto lemma dord_fun_refl: assumes "ord s s" shows "dord_fun ord s s" using assms unfolding dord_fun_def by simp lemma dord_fun_antisym: assumes ord_antisym: "ord s t \ ord t s \ s = t" and "dord_fun ord s t" and "dord_fun ord t s" shows "s = t" proof - from assms(3) have ts: "deg_fun t < deg_fun s \ (deg_fun t = deg_fun s \ ord t s)" unfolding dord_fun_def Let_def . from assms(2) have st: "deg_fun s < deg_fun t \ (deg_fun s = deg_fun t \ ord s t)" unfolding dord_fun_def Let_def . thus ?thesis proof assume "deg_fun s < deg_fun t" thus ?thesis using ts by auto next assume "deg_fun s = deg_fun t \ ord s t" hence "deg_fun s = deg_fun t" and "ord s t" by simp_all from \deg_fun s = deg_fun t\ ts have "ord t s" by simp with \ord s t\ show ?thesis by (rule ord_antisym) qed qed lemma dord_fun_trans: assumes ord_trans: "ord s t \ ord t u \ ord s u" and "dord_fun ord s t" and "dord_fun ord t u" shows "dord_fun ord s u" proof - from assms(3) have ts: "deg_fun t < deg_fun u \ (deg_fun t = deg_fun u \ ord t u)" unfolding dord_fun_def Let_def . from assms(2) have st: "deg_fun s < deg_fun t \ (deg_fun s = deg_fun t \ ord s t)" unfolding dord_fun_def Let_def . thus ?thesis proof assume "deg_fun s < deg_fun t" from this dord_fun_degD[OF assms(3)] have "deg_fun s < deg_fun u" by simp thus ?thesis by (simp add: dord_fun_def Let_def) next assume "deg_fun s = deg_fun t \ ord s t" hence "deg_fun s = deg_fun t" and "ord s t" by simp_all from ts show ?thesis proof assume "deg_fun t < deg_fun u" hence "deg_fun s < deg_fun u" using \deg_fun s = deg_fun t\ by simp thus ?thesis by (simp add: dord_fun_def Let_def) next assume "deg_fun t = deg_fun u \ ord t u" hence "deg_fun t = deg_fun u" and "ord t u" by simp_all from ord_trans[OF \ord s t\ \ord t u\] \deg_fun s = deg_fun t\ \deg_fun t = deg_fun u\ show ?thesis by (simp add: dord_fun_def Let_def) qed qed qed lemma dord_fun_lin: "dord_fun ord s t \ dord_fun ord t s" if "ord s t \ ord t s" for s t::"'a \ 'b::{ordered_comm_monoid_add, linorder}" proof (intro disjCI) assume "\ dord_fun ord t s" hence "deg_fun s \ deg_fun t \ (deg_fun t \ deg_fun s \ \ ord t s)" unfolding dord_fun_def Let_def by auto hence "deg_fun s \ deg_fun t" and dis1: "deg_fun t \ deg_fun s \ \ ord t s" by simp_all show "dord_fun ord s t" unfolding dord_fun_def Let_def proof (intro disjCI) assume "\ (deg_fun s = deg_fun t \ ord s t)" hence dis2: "deg_fun s \ deg_fun t \ \ ord s t" by simp show "deg_fun s < deg_fun t" proof (cases "deg_fun s = deg_fun t") case True from True dis1 have "\ ord t s" by simp from True dis2 have "\ ord s t" by simp from \\ ord s t\ \\ ord t s\ that show ?thesis by simp next case False from this \deg_fun s \ deg_fun t\ show ?thesis by simp qed qed qed lemma dord_fun_zero_min: fixes s t::"'a \ 'b::add_linorder_min" assumes ord_refl: "\t. ord t t" and "finite (supp_fun s)" shows "dord_fun ord 0 s" unfolding dord_fun_def Let_def deg_fun_zero proof (rule disjCI) assume "\ (0 = deg_fun s \ ord 0 s)" hence dis: "deg_fun s \ 0 \ \ ord 0 s" by simp show "0 < deg_fun s" proof (cases "deg_fun s = 0") case True hence "s = 0" using deg_fun_eq_0_iff[OF assms(2)] by auto hence "ord 0 s" using ord_refl by simp with True dis show ?thesis by simp next case False thus ?thesis by (auto simp: zero_less_iff_neq_zero) qed qed lemma dord_fun_plus_monotone: fixes s t u ::"'a \ 'b::{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le}" assumes ord_monotone: "ord s t \ ord (s + u) (t + u)" and "finite (supp_fun s)" and "finite (supp_fun t)" and "finite (supp_fun u)" and "dord_fun ord s t" shows "dord_fun ord (s + u) (t + u)" proof - from assms(5) have "deg_fun s < deg_fun t \ (deg_fun s = deg_fun t \ ord s t)" unfolding dord_fun_def Let_def . thus ?thesis proof assume "deg_fun s < deg_fun t" hence "deg_fun (s + u) < deg_fun (t + u)" by (auto simp: deg_fun_plus[OF _ assms(4)] assms(2) assms(3)) thus ?thesis unfolding dord_fun_def Let_def by simp next assume "deg_fun s = deg_fun t \ ord s t" hence "deg_fun s = deg_fun t" and "ord s t" by simp_all from \deg_fun s = deg_fun t\ have "deg_fun (s + u) = deg_fun (t + u)" by (auto simp: deg_fun_plus[OF _ assms(4)] assms(2) assms(3)) from this ord_monotone[OF \ord s t\] show ?thesis unfolding dord_fun_def Let_def by simp qed qed end (* linorder *) context wellorder begin subsubsection \Degree-Lexicographic Term Order\ definition dlex_fun::"('a \ 'b::ordered_comm_monoid_add) \ ('a \ 'b) \ bool" where "dlex_fun \ dord_fun lex_fun" definition "dlex_fun_strict s t \ dlex_fun s t \ \ dlex_fun t s" lemma dlex_fun_refl: shows "dlex_fun s s" unfolding dlex_fun_def by (rule dord_fun_refl, rule lex_fun_refl) lemma dlex_fun_antisym: assumes "dlex_fun s t" and "dlex_fun t s" shows "s = t" by (rule dord_fun_antisym, erule lex_fun_antisym, assumption, simp_all only: dlex_fun_def[symmetric], fact+) lemma dlex_fun_trans: assumes "dlex_fun s t" and "dlex_fun t u" shows "dlex_fun s u" by (simp only: dlex_fun_def, rule dord_fun_trans, erule lex_fun_trans, assumption, simp_all only: dlex_fun_def[symmetric], fact+) lemma dlex_fun_lin: "dlex_fun s t \ dlex_fun t s" for s t::"('a \ 'b::{ordered_comm_monoid_add, linorder})" unfolding dlex_fun_def by (rule dord_fun_lin, rule lex_fun_lin) corollary dlex_fun_strict_alt [code]: "dlex_fun_strict s t = (\ dlex_fun t s)" for s t::"'a \ 'b::{ordered_comm_monoid_add, linorder}" unfolding dlex_fun_strict_def using dlex_fun_lin by auto lemma dlex_fun_zero_min: fixes s t::"('a \ 'b::add_linorder_min)" assumes "finite (supp_fun s)" shows "dlex_fun 0 s" unfolding dlex_fun_def by (rule dord_fun_zero_min, rule lex_fun_refl, fact) lemma dlex_fun_plus_monotone: fixes s t u::"'a \ 'b::{ordered_cancel_comm_monoid_add, ordered_ab_semigroup_add_imp_le}" assumes "finite (supp_fun s)" and "finite (supp_fun t)" and "finite (supp_fun u)" and "dlex_fun s t" shows "dlex_fun (s + u) (t + u)" using lex_fun_plus_monotone[of s t u] assms unfolding dlex_fun_def by (rule dord_fun_plus_monotone) subsubsection \Degree-Reverse-Lexicographic Term Order\ abbreviation rlex_fun::"('a \ 'b) \ ('a \ 'b::order) \ bool" where "rlex_fun s t \ lex_fun t s" text \Note that @{const rlex_fun} is not precisely the reverse-lexicographic order relation on power-products. Normally, the @{emph \last\} (i.\,e. highest) indeterminate whose exponent differs in the two power-products to be compared is taken, but since we do not require the domain to be finite, there might not be such a last indeterminate. Therefore, we simply take the converse of @{const lex_fun}.\ definition drlex_fun::"('a \ 'b::ordered_comm_monoid_add) \ ('a \ 'b) \ bool" where "drlex_fun \ dord_fun rlex_fun" definition "drlex_fun_strict s t \ drlex_fun s t \ \ drlex_fun t s" lemma drlex_fun_refl: shows "drlex_fun s s" unfolding drlex_fun_def by (rule dord_fun_refl, fact lex_fun_refl) lemma drlex_fun_antisym: assumes "drlex_fun s t" and "drlex_fun t s" shows "s = t" by (rule dord_fun_antisym, erule lex_fun_antisym, assumption, simp_all only: drlex_fun_def[symmetric], fact+) lemma drlex_fun_trans: assumes "drlex_fun s t" and "drlex_fun t u" shows "drlex_fun s u" by (simp only: drlex_fun_def, rule dord_fun_trans, erule lex_fun_trans, assumption, simp_all only: drlex_fun_def[symmetric], fact+) lemma drlex_fun_lin: "drlex_fun s t \ drlex_fun t s" for s t::"('a \ 'b::{ordered_comm_monoid_add, linorder})" unfolding drlex_fun_def by (rule dord_fun_lin, rule lex_fun_lin) corollary drlex_fun_strict_alt [code]: "drlex_fun_strict s t = (\ drlex_fun t s)" for s t::"'a \ 'b::{ordered_comm_monoid_add, linorder}" unfolding drlex_fun_strict_def using drlex_fun_lin by auto lemma drlex_fun_zero_min: fixes s t::"('a \ 'b::add_linorder_min)" assumes "finite (supp_fun s)" shows "drlex_fun 0 s" unfolding drlex_fun_def by (rule dord_fun_zero_min, rule lex_fun_refl, fact) lemma drlex_fun_plus_monotone: fixes s t u::"'a \ 'b::{ordered_cancel_comm_monoid_add, ordered_ab_semigroup_add_imp_le}" assumes "finite (supp_fun s)" and "finite (supp_fun t)" and "finite (supp_fun u)" and "drlex_fun s t" shows "drlex_fun (s + u) (t + u)" using lex_fun_plus_monotone[of t s u] assms unfolding drlex_fun_def by (rule dord_fun_plus_monotone) end (* wellorder *) text\Every finite linear ordering is also a well-ordering. This fact is particularly useful when working with fixed finite sets of indeterminates.\ class finite_linorder = finite + linorder begin subclass wellorder proof fix P::"'a \ bool" and a assume hyp: "\x. (\y. (y < x) \ P y) \ P x" show "P a" proof (rule ccontr) assume "\ P a" have "finite {x. \ P x}" (is "finite ?A") by simp from \\ P a\ have "a \ ?A" by simp hence "?A \ {}" by auto from ex_min[OF \finite ?A\ this] obtain b where "b \ ?A" and b_min: "\y\?A. b \ y" by auto from \b \ ?A\ have "\ P b" by simp with hyp[of b] obtain y where "y < b" and "\ P y" by auto from \\ P y\ have "y \ ?A" by simp with b_min have "b \ y" by simp with \y < b\ show False by simp qed qed end subsection \Type @{type poly_mapping}\ lemma poly_mapping_eq_zeroI: assumes "keys s = {}" shows "s = (0::('a, 'b::zero) poly_mapping)" proof (rule poly_mapping_eqI, simp) fix x from assms show "lookup s x = 0" by auto qed lemma keys_plus_ninv_comm_monoid_add: "keys (s + t) = keys s \ keys (t::'a \\<^sub>0 'b::ninv_comm_monoid_add)" proof (rule, fact Poly_Mapping.keys_add, rule) fix x assume "x \ keys s \ keys t" thus "x \ keys (s + t)" proof assume "x \ keys s" thus ?thesis by (metis in_keys_iff lookup_add plus_eq_zero) next assume "x \ keys t" thus ?thesis by (metis in_keys_iff lookup_add plus_eq_zero_2) qed qed lemma lookup_zero_fun: "lookup 0 = 0" by (simp only: zero_poly_mapping.rep_eq zero_fun_def) lemma lookup_plus_fun: "lookup (s + t) = lookup s + lookup t" by (simp only: plus_poly_mapping.rep_eq plus_fun_def) lemma lookup_uminus_fun: "lookup (- s) = - lookup s" by (fact uminus_poly_mapping.rep_eq) lemma lookup_minus_fun: "lookup (s - t) = lookup s - lookup t" by (simp only: minus_poly_mapping.rep_eq, rule, simp only: minus_apply) lemma poly_mapping_adds_iff: "s adds t \ lookup s adds lookup t" unfolding adds_def proof assume "\k. t = s + k" then obtain k where *: "t = s + k" .. show "\k. lookup t = lookup s + k" proof from * show "lookup t = lookup s + lookup k" by (simp only: lookup_plus_fun) qed next assume "\k. lookup t = lookup s + k" then obtain k where *: "lookup t = lookup s + k" .. have **: "k \ {f. finite {x. f x \ 0}}" proof have "finite {x. lookup t x \ 0}" by transfer hence "finite {x. lookup s x + k x \ 0}" by (simp only: * plus_fun_def) moreover have "finite {x. lookup s x \ 0}" by transfer ultimately show "finite {x. k x \ 0}" by (rule finite_neq_0_inv', simp) qed show "\k. t = s + k" proof show "t = s + Abs_poly_mapping k" by (rule poly_mapping_eqI, simp add: * lookup_add Abs_poly_mapping_inverse[OF **]) qed qed subsubsection \@{typ "('a, 'b) poly_mapping"} belongs to class @{class comm_powerprod}\ instance poly_mapping :: (type, cancel_comm_monoid_add) comm_powerprod by standard subsubsection \@{typ "('a, 'b) poly_mapping"} belongs to class @{class ninv_comm_monoid_add}\ instance poly_mapping :: (type, ninv_comm_monoid_add) ninv_comm_monoid_add proof (standard, transfer) fix s t::"'a \ 'b" assume "(\k. s k + t k) = (\_. 0)" hence "s + t = 0" by (simp only: plus_fun_def zero_fun_def) hence "s = 0" by (rule plus_eq_zero) thus "s = (\_. 0)" by (simp only: zero_fun_def) qed subsubsection \@{typ "('a, 'b) poly_mapping"} belongs to class @{class lcs_powerprod}\ instantiation poly_mapping :: (type, add_linorder) lcs_powerprod begin lift_definition lcs_poly_mapping::"('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b)" is "\s t. \x. max (s x) (t x)" proof - fix fun1 fun2::"'a \ 'b" assume "finite {t. fun1 t \ 0}" and "finite {t. fun2 t \ 0}" from finite_neq_0'[OF this, of max] show "finite {t. max (fun1 t) (fun2 t) \ 0}" by (auto simp: max_def) qed lemma adds_poly_mappingI: assumes "lookup s \ lookup (t::'a \\<^sub>0 'b)" shows "s adds t" unfolding poly_mapping_adds_iff using assms by (rule adds_funI) lemma lookup_lcs_fun: "lookup (lcs s t) = lcs (lookup s) (lookup (t:: 'a \\<^sub>0 'b))" by (simp only: lcs_poly_mapping.rep_eq lcs_fun_def) instance by (standard, simp_all only: poly_mapping_adds_iff lookup_lcs_fun, rule adds_lcs, elim lcs_adds, assumption, rule poly_mapping_eqI, simp only: lookup_lcs_fun lcs_comm) end lemma adds_poly_mapping: "s adds t \ lookup s \ lookup t" for s t::"'a \\<^sub>0 'b::add_linorder_min" by (simp only: poly_mapping_adds_iff adds_fun) lemma lookup_gcs_fun: "lookup (gcs s (t::'a \\<^sub>0 ('b::add_linorder))) = gcs (lookup s) (lookup t)" proof fix x show "lookup (gcs s t) x = gcs (lookup s) (lookup t) x" by (simp add: gcs_def lookup_minus lookup_add lookup_lcs_fun) qed subsubsection \@{typ "('a, 'b) poly_mapping"} belongs to class @{class ulcs_powerprod}\ instance poly_mapping :: (type, add_linorder_min) ulcs_powerprod .. subsubsection \Power-products in a given set of indeterminates.\ lemma adds_except: "s adds t = (except s V adds except t V \ except s (- V) adds except t (- V))" for s t :: "'a \\<^sub>0 'b::add_linorder" by (simp add: poly_mapping_adds_iff adds_except_fun[of "lookup s", where V=V] except.rep_eq) lemma adds_except_singleton: "s adds t \ (except s {v} adds except t {v} \ lookup s v adds lookup t v)" for s t :: "'a \\<^sub>0 'b::add_linorder" by (simp add: poly_mapping_adds_iff adds_except_fun_singleton[of "lookup s", where v=v] except.rep_eq) subsubsection \Dickson's lemma for power-products in finitely many indeterminates\ context countable begin definition elem_index :: "'a \ nat" where "elem_index = (SOME f. inj f)" lemma inj_elem_index: "inj elem_index" unfolding elem_index_def using ex_inj by (rule someI_ex) lemma elem_index_inj: assumes "elem_index x = elem_index y" shows "x = y" using inj_elem_index assms by (rule injD) lemma finite_nat_seg: "finite {x. elem_index x < n}" proof (rule finite_imageD) have "elem_index ` {x. elem_index x < n} \ {0..\<^sub>0 'b::add_wellorder. keys x \ V}" proof (rule almost_full_onI) fix seq::"nat \ 'a \\<^sub>0 'b" assume a: "\i. seq i \ {x::'a \\<^sub>0 'b. keys x \ V}" define seq' where "seq' = (\i. lookup (seq i))" from assms have "almost_full_on (adds) {x::'a \ 'b. supp_fun x \ V}" by (rule Dickson_fun) moreover from a have "\i. seq' i \ {x::'a \ 'b. supp_fun x \ V}" by (auto simp: seq'_def keys_eq_supp) ultimately obtain i j where "i < j" and "seq' i adds seq' j" by (rule almost_full_onD) from this(2) have "seq i adds seq j" by (simp add: seq'_def poly_mapping_adds_iff) with \i < j\ show "good (adds) seq" by (rule goodI) qed definition varnum :: "'x set \ ('x::countable \\<^sub>0 'b::zero) \ nat" where "varnum X t = (if keys t - X = {} then 0 else Suc (Max (elem_index ` (keys t - X))))" lemma elem_index_less_varnum: assumes "x \ keys t" obtains "x \ X" | "elem_index x < varnum X t" proof (cases "x \ X") case True thus ?thesis .. next case False with assms have 1: "x \ keys t - X" by simp hence "keys t - X \ {}" by blast hence eq: "varnum X t = Suc (Max (elem_index ` (keys t - X)))" by (simp add: varnum_def) hence "elem_index x < varnum X t" using 1 by (simp add: less_Suc_eq_le) thus ?thesis .. qed lemma varnum_plus: "varnum X (s + t) = max (varnum X s) (varnum X (t::'x::countable \\<^sub>0 'b::ninv_comm_monoid_add))" proof (simp add: varnum_def keys_plus_ninv_comm_monoid_add image_Un Un_Diff del: Diff_eq_empty_iff, intro impI) assume 1: "keys s - X \ {}" and 2: "keys t - X \ {}" have "finite (elem_index ` (keys s - X))" by simp moreover from 1 have "elem_index ` (keys s - X) \ {}" by simp moreover have "finite (elem_index ` (keys t - X))" by simp moreover from 2 have "elem_index ` (keys t - X) \ {}" by simp ultimately show "Max (elem_index ` (keys s - X) \ elem_index ` (keys t - X)) = max (Max (elem_index ` (keys s - X))) (Max (elem_index ` (keys t - X)))" by (rule Max_Un) qed lemma dickson_grading_varnum: assumes "finite X" shows "dickson_grading ((varnum X)::('x::countable \\<^sub>0 'b::add_wellorder) \ nat)" using varnum_plus proof (rule dickson_gradingI) fix m::nat let ?V = "X \ {x. elem_index x < m}" have "{t::'x \\<^sub>0 'b. varnum X t \ m} \ {t. keys t \ ?V}" proof (rule, simp, intro subsetI, simp) fix t::"'x \\<^sub>0 'b" and x::'x assume "varnum X t \ m" assume "x \ keys t" thus "x \ X \ elem_index x < m" proof (rule elem_index_less_varnum) assume "x \ X" thus ?thesis .. next assume "elem_index x < varnum X t" hence "elem_index x < m" using \varnum X t \ m\ by (rule less_le_trans) thus ?thesis .. qed qed thus "almost_full_on (adds) {t::'x \\<^sub>0 'b. varnum X t \ m}" proof (rule almost_full_on_subset) from assms finite_nat_seg have "finite ?V" by (rule finite_UnI) thus "almost_full_on (adds) {t::'x \\<^sub>0 'b. keys t \ ?V}" by (rule Dickson_poly_mapping) qed qed corollary dickson_grading_varnum_empty: "dickson_grading ((varnum {})::(_ \\<^sub>0 _::add_wellorder) \ nat)" using finite.emptyI by (rule dickson_grading_varnum) lemma varnum_le_iff: "varnum X t \ n \ keys t \ X \ {x. elem_index x < n}" by (auto simp: varnum_def Suc_le_eq) lemma varnum_zero [simp]: "varnum X 0 = 0" by (simp add: varnum_def) lemma varnum_empty_eq_zero_iff: "varnum {} t = 0 \ t = 0" proof assume "varnum {} t = 0" hence "keys t = {}" by (simp add: varnum_def split: if_splits) thus "t = 0" by (rule poly_mapping_eq_zeroI) qed simp instance poly_mapping :: (countable, add_wellorder) graded_dickson_powerprod by standard (rule, fact dickson_grading_varnum_empty) instance poly_mapping :: (finite, add_wellorder) dickson_powerprod proof have "finite (UNIV::'a set)" by simp hence "almost_full_on (adds) {x::'a \\<^sub>0 'b. keys x \ UNIV}" by (rule Dickson_poly_mapping) thus "almost_full_on (adds) (UNIV::('a \\<^sub>0 'b) set)" by simp qed subsubsection \Lexicographic Term Order\ definition lex_pm :: "('a \\<^sub>0 'b) \ ('a::linorder \\<^sub>0 'b::{zero,linorder}) \ bool" where "lex_pm = (\)" definition lex_pm_strict :: "('a \\<^sub>0 'b) \ ('a::linorder \\<^sub>0 'b::{zero,linorder}) \ bool" where "lex_pm_strict = (<)" lemma lex_pm_alt: "lex_pm s t = (s = t \ (\x. lookup s x < lookup t x \ (\y lex_pm t s \ s = t" by (simp add: lex_pm_def) lemma lex_pm_trans: "lex_pm s t \ lex_pm t u \ lex_pm s u" by (simp add: lex_pm_def) lemma lex_pm_lin: "lex_pm s t \ lex_pm t s" by (simp add: lex_pm_def linear) corollary lex_pm_strict_alt [code]: "lex_pm_strict s t = (\ lex_pm t s)" by (auto simp: lex_pm_strict_def lex_pm_def) lemma lex_pm_zero_min: "lex_pm 0 s" for s::"_ \\<^sub>0 _::add_linorder_min" proof (rule ccontr) assume "\ lex_pm 0 s" hence "lex_pm_strict s 0" by (simp add: lex_pm_strict_alt) thus False by (simp add: lex_pm_strict_def less_poly_mapping.rep_eq less_fun_def) qed lemma lex_pm_plus_monotone: "lex_pm s t \ lex_pm (s + u) (t + u)" for s t::"_ \\<^sub>0 _::{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le}" by (simp add: lex_pm_def add_right_mono) subsubsection \Degree\ lift_definition deg_pm::"('a \\<^sub>0 'b::comm_monoid_add) \ 'b" is deg_fun . lemma deg_pm_zero[simp]: "deg_pm 0 = 0" by (simp add: deg_pm.rep_eq lookup_zero_fun) lemma deg_pm_eq_0_iff[simp]: "deg_pm s = 0 \ s = 0" for s::"'a \\<^sub>0 'b::add_linorder_min" by (simp only: deg_pm.rep_eq poly_mapping_eq_iff lookup_zero_fun, rule deg_fun_eq_0_iff, simp add: keys_eq_supp[symmetric]) lemma deg_pm_superset: assumes "keys s \ A" and "finite A" shows "deg_pm s = (\x\A. lookup s x)" using assms by (simp only: deg_pm.rep_eq keys_eq_supp, elim deg_fun_superset) lemma deg_pm_plus: "deg_pm (s + t) = deg_pm s + deg_pm (t::'a \\<^sub>0 'b::comm_monoid_add)" by (simp only: deg_pm.rep_eq lookup_plus_fun, rule deg_fun_plus, simp_all add: keys_eq_supp[symmetric]) lemma deg_pm_single: "deg_pm (Poly_Mapping.single x k) = k" proof - have "keys (Poly_Mapping.single x k) \ {x}" by simp moreover have "finite {x}" by simp ultimately have "deg_pm (Poly_Mapping.single x k) = (\y\{x}. lookup (Poly_Mapping.single x k) y)" by (rule deg_pm_superset) also have "... = k" by simp finally show ?thesis . qed subsubsection \General Degree-Orders\ context linorder begin lift_definition dord_pm::"(('a \\<^sub>0 'b::ordered_comm_monoid_add) \ ('a \\<^sub>0 'b) \ bool) \ ('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b) \ bool" is dord_fun by (metis local.dord_fun_def) lemma dord_pm_alt: "dord_pm ord = (\x y. deg_pm x < deg_pm y \ (deg_pm x = deg_pm y \ ord x y))" by (intro ext) (transfer, simp add: dord_fun_def Let_def) lemma dord_pm_degD: assumes "dord_pm ord s t" shows "deg_pm s \ deg_pm t" using assms by (simp only: dord_pm.rep_eq deg_pm.rep_eq, elim dord_fun_degD) lemma dord_pm_refl: assumes "ord s s" shows "dord_pm ord s s" using assms by (simp only: dord_pm.rep_eq, intro dord_fun_refl, simp add: lookup_inverse) lemma dord_pm_antisym: assumes "ord s t \ ord t s \ s = t" and "dord_pm ord s t" and "dord_pm ord t s" shows "s = t" using assms proof (simp only: dord_pm.rep_eq poly_mapping_eq_iff) assume 1: "(ord s t \ ord t s \ lookup s = lookup t)" assume 2: "dord_fun (map_fun Abs_poly_mapping id \ ord \ Abs_poly_mapping) (lookup s) (lookup t)" assume 3: "dord_fun (map_fun Abs_poly_mapping id \ ord \ Abs_poly_mapping) (lookup t) (lookup s)" from _ 2 3 show "lookup s = lookup t" by (rule dord_fun_antisym, simp add: lookup_inverse 1) qed lemma dord_pm_trans: assumes "ord s t \ ord t u \ ord s u" and "dord_pm ord s t" and "dord_pm ord t u" shows "dord_pm ord s u" using assms proof (simp only: dord_pm.rep_eq poly_mapping_eq_iff) assume 1: "(ord s t \ ord t u \ ord s u)" assume 2: "dord_fun (map_fun Abs_poly_mapping id \ ord \ Abs_poly_mapping) (lookup s) (lookup t)" assume 3: "dord_fun (map_fun Abs_poly_mapping id \ ord \ Abs_poly_mapping) (lookup t) (lookup u)" from _ 2 3 show "dord_fun (map_fun Abs_poly_mapping id \ ord \ Abs_poly_mapping) (lookup s) (lookup u)" by (rule dord_fun_trans, simp add: lookup_inverse 1) qed lemma dord_pm_lin: "dord_pm ord s t \ dord_pm ord t s" if "ord s t \ ord t s" for s t::"'a \\<^sub>0 'b::{ordered_comm_monoid_add, linorder}" using that by (simp only: dord_pm.rep_eq, intro dord_fun_lin, simp add: lookup_inverse) lemma dord_pm_zero_min: "dord_pm ord 0 s" if ord_refl: "\t. ord t t" for s t::"'a \\<^sub>0 'b::add_linorder_min" using that by (simp only: dord_pm.rep_eq lookup_zero_fun, intro dord_fun_zero_min, simp add: lookup_inverse, simp add: keys_eq_supp[symmetric]) lemma dord_pm_plus_monotone: fixes s t u ::"'a \\<^sub>0 'b::{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le}" assumes "ord s t \ ord (s + u) (t + u)" and "dord_pm ord s t" shows "dord_pm ord (s + u) (t + u)" using assms by (simp only: dord_pm.rep_eq lookup_plus_fun, intro dord_fun_plus_monotone, simp add: lookup_inverse lookup_plus_fun[symmetric], simp add: keys_eq_supp[symmetric], simp add: keys_eq_supp[symmetric], simp add: keys_eq_supp[symmetric], simp add: lookup_inverse) end (* linorder *) subsubsection \Degree-Lexicographic Term Order\ definition dlex_pm::"('a::linorder \\<^sub>0 'b::{ordered_comm_monoid_add,linorder}) \ ('a \\<^sub>0 'b) \ bool" where "dlex_pm \ dord_pm lex_pm" definition "dlex_pm_strict s t \ dlex_pm s t \ \ dlex_pm t s" lemma dlex_pm_refl: "dlex_pm s s" unfolding dlex_pm_def using lex_pm_refl by (rule dord_pm_refl) lemma dlex_pm_antisym: "dlex_pm s t \ dlex_pm t s \ s = t" unfolding dlex_pm_def using lex_pm_antisym by (rule dord_pm_antisym) lemma dlex_pm_trans: "dlex_pm s t \ dlex_pm t u \ dlex_pm s u" unfolding dlex_pm_def using lex_pm_trans by (rule dord_pm_trans) lemma dlex_pm_lin: "dlex_pm s t \ dlex_pm t s" unfolding dlex_pm_def using lex_pm_lin by (rule dord_pm_lin) corollary dlex_pm_strict_alt [code]: "dlex_pm_strict s t = (\ dlex_pm t s)" unfolding dlex_pm_strict_def using dlex_pm_lin by auto lemma dlex_pm_zero_min: "dlex_pm 0 s" for s t::"(_ \\<^sub>0 _::add_linorder_min)" unfolding dlex_pm_def using lex_pm_refl by (rule dord_pm_zero_min) lemma dlex_pm_plus_monotone: "dlex_pm s t \ dlex_pm (s + u) (t + u)" for s t::"_ \\<^sub>0 _::{ordered_ab_semigroup_add_imp_le, ordered_cancel_comm_monoid_add}" unfolding dlex_pm_def using lex_pm_plus_monotone by (rule dord_pm_plus_monotone) subsubsection \Degree-Reverse-Lexicographic Term Order\ definition drlex_pm::"('a::linorder \\<^sub>0 'b::{ordered_comm_monoid_add,linorder}) \ ('a \\<^sub>0 'b) \ bool" where "drlex_pm \ dord_pm (\s t. lex_pm t s)" definition "drlex_pm_strict s t \ drlex_pm s t \ \ drlex_pm t s" lemma drlex_pm_refl: "drlex_pm s s" unfolding drlex_pm_def using lex_pm_refl by (rule dord_pm_refl) lemma drlex_pm_antisym: "drlex_pm s t \ drlex_pm t s \ s = t" unfolding drlex_pm_def using lex_pm_antisym by (rule dord_pm_antisym) lemma drlex_pm_trans: "drlex_pm s t \ drlex_pm t u \ drlex_pm s u" unfolding drlex_pm_def using lex_pm_trans by (rule dord_pm_trans) lemma drlex_pm_lin: "drlex_pm s t \ drlex_pm t s" unfolding drlex_pm_def using lex_pm_lin by (rule dord_pm_lin) corollary drlex_pm_strict_alt [code]: "drlex_pm_strict s t = (\ drlex_pm t s)" unfolding drlex_pm_strict_def using drlex_pm_lin by auto lemma drlex_pm_zero_min: "drlex_pm 0 s" for s t::"(_ \\<^sub>0 _::add_linorder_min)" unfolding drlex_pm_def using lex_pm_refl by (rule dord_pm_zero_min) lemma drlex_pm_plus_monotone: "drlex_pm s t \ drlex_pm (s + u) (t + u)" for s t::"_ \\<^sub>0 _::{ordered_ab_semigroup_add_imp_le, ordered_cancel_comm_monoid_add}" unfolding drlex_pm_def using lex_pm_plus_monotone by (rule dord_pm_plus_monotone) end (* theory *) diff --git a/thys/Types_To_Sets_Extension/Examples/SML_Relativization/Simple_Orders/SML_Simple_Orders.thy b/thys/Types_To_Sets_Extension/Examples/SML_Relativization/Simple_Orders/SML_Simple_Orders.thy --- a/thys/Types_To_Sets_Extension/Examples/SML_Relativization/Simple_Orders/SML_Simple_Orders.thy +++ b/thys/Types_To_Sets_Extension/Examples/SML_Relativization/Simple_Orders/SML_Simple_Orders.thy @@ -1,1708 +1,1696 @@ (* Title: Examples/SML_Relativization/Simple_Orders/SML_Simple_Orders.thy Author: Mihails Milehins Copyright 2021 (C) Mihails Milehins *) section\Relativization of the results about orders\ theory SML_Simple_Orders imports "../../Introduction" "../Foundations/SML_Relations" Complex_Main begin subsection\Class \<^class>\ord\\ subsubsection\Definitions and common properties\ locale ord_ow = fixes U :: "'ao set" and le :: "['ao, 'ao] \ bool" (infix \\\<^sub>o\<^sub>w\ 50) and ls :: "['ao, 'ao] \ bool" (infix \<\<^sub>o\<^sub>w\ 50) begin notation le (\'(\\<^sub>o\<^sub>w')\) and le (infix \\\<^sub>o\<^sub>w\ 50) and ls (\'(<\<^sub>o\<^sub>w')\) and ls (infix \<\<^sub>o\<^sub>w\ 50) abbreviation (input) ge (infix \\\<^sub>o\<^sub>w\ 50) where "x \\<^sub>o\<^sub>w y \ y \\<^sub>o\<^sub>w x" abbreviation (input) gt (infix \>\<^sub>o\<^sub>w\ 50) where "x >\<^sub>o\<^sub>w y \ y <\<^sub>o\<^sub>w x" notation ge (\'(\\<^sub>o\<^sub>w')\) and ge (infix \\\<^sub>o\<^sub>w\ 50) and gt (\'(>\<^sub>o\<^sub>w')\) and gt (infix \>\<^sub>o\<^sub>w\ 50) tts_register_sbts \(\\<^sub>o\<^sub>w)\ | U proof- assume "Domainp AOA = (\x. x \ U)" "bi_unique AOA" "right_total AOA" from tts_AA_eq_transfer[OF this] show ?thesis by auto qed tts_register_sbts \(<\<^sub>o\<^sub>w)\ | U proof- assume "Domainp AOA = (\x. x \ U)" "bi_unique AOA" "right_total AOA" from tts_AA_eq_transfer[OF this] show ?thesis by auto qed end locale ord_pair_ow = ord\<^sub>1: ord_ow U\<^sub>1 le\<^sub>1 ls\<^sub>1 + ord\<^sub>2: ord_ow U\<^sub>2 le\<^sub>2 ls\<^sub>2 for U\<^sub>1 :: "'ao set" and le\<^sub>1 ls\<^sub>1 and U\<^sub>2 :: "'bo set" and le\<^sub>2 ls\<^sub>2 begin notation le\<^sub>1 (\'(\\<^sub>o\<^sub>w\<^sub>.\<^sub>1')\) and le\<^sub>1 (infix \\\<^sub>o\<^sub>w\<^sub>.\<^sub>1\ 50) and ls\<^sub>1 (\'(<\<^sub>o\<^sub>w\<^sub>.\<^sub>1')\) and ls\<^sub>1 (infix \<\<^sub>o\<^sub>w\<^sub>.\<^sub>1\ 50) and le\<^sub>2 (\'(\\<^sub>o\<^sub>w\<^sub>.\<^sub>2')\) and le\<^sub>2 (infix \\\<^sub>o\<^sub>w\<^sub>.\<^sub>2\ 50) and ls\<^sub>2 (\'(<\<^sub>o\<^sub>w\<^sub>.\<^sub>2')\) and ls\<^sub>2 (infix \<\<^sub>o\<^sub>w\<^sub>.\<^sub>2\ 50) notation ord\<^sub>1.ge (\'(\\<^sub>o\<^sub>w\<^sub>.\<^sub>1')\) and ord\<^sub>1.ge (infix \\\<^sub>o\<^sub>w\<^sub>.\<^sub>1\ 50) and ord\<^sub>1.gt (\'(>\<^sub>o\<^sub>w\<^sub>.\<^sub>1')\) and ord\<^sub>1.gt (infix \>\<^sub>o\<^sub>w\<^sub>.\<^sub>1\ 50) and ord\<^sub>2.ge (\'(\\<^sub>o\<^sub>w\<^sub>.\<^sub>2')\) and ord\<^sub>2.ge (infix \\\<^sub>o\<^sub>w\<^sub>.\<^sub>2\ 50) and ord\<^sub>2.gt (\'(>\<^sub>o\<^sub>w\<^sub>.\<^sub>2')\) and ord\<^sub>2.gt (infix \>\<^sub>o\<^sub>w\<^sub>.\<^sub>2\ 50) end ud \ord.lessThan\ (\(with _ : ({..<_}))\ [1000] 10) ud lessThan' \lessThan\ ud \ord.atMost\ (\(with _ : ({.._}))\ [1000] 10) ud atMost' \atMost\ ud \ord.greaterThan\ (\(with _ : ({_<..}))\ [1000] 10) ud greaterThan' \greaterThan\ ud \ord.atLeast\ (\(with _ : ({_..}))\ [1000] 10) ud atLeast' \atLeast\ ud \ord.greaterThanLessThan\ (\(with _ : ({_<..<_}))\ [1000, 999, 1000] 10) ud greaterThanLessThan' \greaterThanLessThan\ ud \ord.atLeastLessThan\ (\(with _ _ : ({_..<_}))\ [1000, 999, 1000, 1000] 10) ud atLeastLessThan' \atLeastLessThan\ ud \ord.greaterThanAtMost\ (\(with _ _ : ({_<.._}))\ [1000, 999, 1000, 999] 10) ud greaterThanAtMost' \greaterThanAtMost\ ud \ord.atLeastAtMost\ (\(with _ : ({_.._}))\ [1000, 1000, 1000] 10) ud atLeastAtMost' \atLeastAtMost\ ud \ord.min\ (\(with _ : \min\ _ _)\ [1000, 1000, 999] 10) ud min' \min\ ud \ord.max\ (\(with _ : \max\ _ _)\ [1000, 1000, 999] 10) ud max' \max\ ctr relativization synthesis ctr_simps assumes [transfer_domain_rule, transfer_rule]: "Domainp A = (\x. x \ U)" and [transfer_rule]: "right_total A" trp (?'a A) in lessThan_ow: lessThan.with_def (\(on _ with _ : ({..<_}))\ [1000, 1000, 1000] 10) and atMost_ow: atMost.with_def (\(on _ with _ : ({.._}))\ [1000, 1000, 1000] 10) and greaterThan_ow: greaterThan.with_def (\(on _ with _: ({_<..}))\ [1000, 1000, 1000] 10) and atLeast_ow: atLeast.with_def (\(on _ with _ : ({_..}))\ [1000, 1000, 1000] 10) ctr relativization synthesis ctr_simps assumes [transfer_domain_rule, transfer_rule]: "Domainp A = (\x. x \ U)" and [transfer_rule]: "bi_unique A" "right_total A" trp (?'a A) in greaterThanLessThan_ow: greaterThanLessThan.with_def (\(on _ with _ : ({_<..<_}))\ [1000, 1000, 1000, 1000] 10) and atLeastLessThan_ow: atLeastLessThan.with_def (\(on _ with _ _ : ({_..<_}))\ [1000, 1000, 999, 1000, 1000] 10) and greaterThanAtMost_ow: greaterThanAtMost.with_def (\(on _ with _ _ : ({_<.._}))\ [1000, 1000, 999, 1000, 1000] 10) and atLeastAtMost_ow: atLeastAtMost.with_def (\(on _ with _ : ({_.._}))\ [1000, 1000, 1000, 1000] 10) ctr parametricity in min_ow: min.with_def and max_ow: max.with_def context ord_ow begin abbreviation lessThan :: "'ao \ 'ao set" ("(1{..<\<^sub>o\<^sub>w_})") where "{..<\<^sub>o\<^sub>w u} \ on U with (<\<^sub>o\<^sub>w) : {.. 'ao set" ("(1{..\<^sub>o\<^sub>w_})") where "{..\<^sub>o\<^sub>w u} \ on U with (\\<^sub>o\<^sub>w) : {..u}" abbreviation greaterThan :: "'ao \ 'ao set" ("(1{_<\<^sub>o\<^sub>w..})") where "{l<\<^sub>o\<^sub>w..} \ on U with (<\<^sub>o\<^sub>w) : {l<..}" abbreviation atLeast :: "'ao \ 'ao set" ("(1{_..\<^sub>o\<^sub>w})") where "atLeast l \ on U with (\\<^sub>o\<^sub>w) : {l..}" abbreviation greaterThanLessThan :: "'ao \ 'ao \ 'ao set" ("(1{_<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>w_})") where "{l<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wu} \ on U with (<\<^sub>o\<^sub>w) : {l<.. 'ao \ 'ao set" ("(1{_..<\<^sub>o\<^sub>w_})") where "{l..<\<^sub>o\<^sub>w u} \ on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {l<..u}" abbreviation greaterThanAtMost :: "'ao \ 'ao \ 'ao set" ("(1{_<\<^sub>o\<^sub>w.._})") where "{l<\<^sub>o\<^sub>w..u} \ on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {l<..u}" abbreviation atLeastAtMost :: "'ao \ 'ao \ 'ao set" ("(1{_..\<^sub>o\<^sub>w_})") where "{l..\<^sub>o\<^sub>wu} \ on U with (\\<^sub>o\<^sub>w) : {l..u}" abbreviation min :: "'ao \ 'ao \ 'ao" where "min \ min.with (\\<^sub>o\<^sub>w)" abbreviation max :: "'ao \ 'ao \ 'ao" where "max \ max.with (\\<^sub>o\<^sub>w)" end context ord_pair_ow begin notation ord\<^sub>1.lessThan ("(1{..<\<^sub>o\<^sub>w\<^sub>.\<^sub>1_})") notation ord\<^sub>1.atMost ("(1{..\<^sub>o\<^sub>w\<^sub>.\<^sub>1_})") notation ord\<^sub>1.greaterThan ("(1{_<\<^sub>o\<^sub>w\<^sub>.\<^sub>1..})") notation ord\<^sub>1.atLeast ("(1{_..\<^sub>o\<^sub>w\<^sub>.\<^sub>1})") notation ord\<^sub>1.greaterThanLessThan ("(1{_<\<^sub>o\<^sub>w\<^sub>.\<^sub>1..<\<^sub>o\<^sub>w\<^sub>.\<^sub>1_})") notation ord\<^sub>1.atLeastLessThan ("(1{_..<\<^sub>o\<^sub>w\<^sub>.\<^sub>1_})") notation ord\<^sub>1.greaterThanAtMost ("(1{_<\<^sub>o\<^sub>w\<^sub>.\<^sub>1.._})") notation ord\<^sub>1.atLeastAtMost ("(1{_..\<^sub>o\<^sub>w\<^sub>.\<^sub>1_})") notation ord\<^sub>2.lessThan ("(1{..<\<^sub>o\<^sub>w\<^sub>.\<^sub>2_})") notation ord\<^sub>2.atMost ("(1{..\<^sub>o\<^sub>w\<^sub>.\<^sub>2_})") notation ord\<^sub>2.greaterThan ("(1{_<\<^sub>o\<^sub>w\<^sub>.\<^sub>2..})") notation ord\<^sub>2.atLeast ("(1{_..\<^sub>o\<^sub>w\<^sub>.\<^sub>2})") notation ord\<^sub>2.greaterThanLessThan ("(1{_<\<^sub>o\<^sub>w\<^sub>.\<^sub>2..<\<^sub>o\<^sub>w\<^sub>.\<^sub>2_})") notation ord\<^sub>2.atLeastLessThan ("(1{_..<\<^sub>o\<^sub>w\<^sub>.\<^sub>2_})") notation ord\<^sub>2.greaterThanAtMost ("(1{_<\<^sub>o\<^sub>w\<^sub>.\<^sub>2.._})") notation ord\<^sub>2.atLeastAtMost ("(1{_..\<^sub>o\<^sub>w\<^sub>.\<^sub>2_})") end subsection\Preorders\ subsubsection\Definitions and common properties\ locale partial_preordering_ow = fixes U :: "'ao set" and le :: "'ao \ 'ao \ bool" (infix "\<^bold>\\<^sub>o\<^sub>w" 50) assumes refl: "a \ U \ a \<^bold>\\<^sub>o\<^sub>w a" and trans: "\ a \ U; b \ U; c \ U; a \<^bold>\\<^sub>o\<^sub>w b; b \<^bold>\\<^sub>o\<^sub>w c \ \ a \<^bold>\\<^sub>o\<^sub>w c" begin notation le (infix "\<^bold>\\<^sub>o\<^sub>w" 50) end locale preordering_ow = partial_preordering_ow U le for U :: "'ao set" and le :: "'ao \ 'ao \ bool" (infix "\<^bold>\\<^sub>o\<^sub>w" 50) + fixes ls :: \'ao \ 'ao \ bool\ (infix "\<^bold><\<^sub>o\<^sub>w" 50) assumes strict_iff_not: "\ a \ U; b \ U \ \ a \<^bold><\<^sub>o\<^sub>w b \ a \<^bold>\\<^sub>o\<^sub>w b \ \ b \<^bold>\\<^sub>o\<^sub>w a" locale preorder_ow = ord_ow U le ls for U :: "'ao set" and le ls + assumes less_le_not_le: "\ x \ U; y \ U \ \ x <\<^sub>o\<^sub>w y \ x \\<^sub>o\<^sub>w y \ \ (y \\<^sub>o\<^sub>w x)" and order_refl[iff]: "x \ U \ x \\<^sub>o\<^sub>w x" and order_trans: "\ x \ U; y \ U; z \ U; x \\<^sub>o\<^sub>w y; y \\<^sub>o\<^sub>w z \ \ x \\<^sub>o\<^sub>w z" begin sublocale order: preordering_ow U \(\\<^sub>o\<^sub>w)\ \(<\<^sub>o\<^sub>w)\ + dual_order: preordering_ow U \(\\<^sub>o\<^sub>w)\ \(>\<^sub>o\<^sub>w)\ apply unfold_locales subgoal by auto subgoal by (meson order_trans) subgoal using less_le_not_le by simp subgoal by (meson order_trans) subgoal by (metis less_le_not_le) done end locale ord_preorder_ow = ord\<^sub>1: ord_ow U\<^sub>1 le\<^sub>1 ls\<^sub>1 + ord\<^sub>2: preorder_ow U\<^sub>2 le\<^sub>2 ls\<^sub>2 for U\<^sub>1 :: "'ao set" and le\<^sub>1 ls\<^sub>1 and U\<^sub>2 :: "'bo set" and le\<^sub>2 ls\<^sub>2 begin sublocale ord_pair_ow . end locale preorder_pair_ow = ord\<^sub>1: preorder_ow U\<^sub>1 le\<^sub>1 ls\<^sub>1 + ord\<^sub>2: preorder_ow U\<^sub>2 le\<^sub>2 ls\<^sub>2 for U\<^sub>1 :: "'ao set" and le\<^sub>1 and ls\<^sub>1 and U\<^sub>2 :: "'bo set" and le\<^sub>2 and ls\<^sub>2 begin sublocale ord_preorder_ow .. end ud \preordering_bdd.bdd\ ud \preorder.bdd_above\ ud bdd_above' \bdd_above\ ud \preorder.bdd_below\ ud bdd_below' \bdd_below\ ctr relativization synthesis ctr_simps assumes [transfer_domain_rule, transfer_rule]: "Domainp A = (\x. x \ U)" and [transfer_rule]: "right_total A" trp (?'a A) in bdd_ow: bdd.with_def (\(on _ with _ : \bdd\ _)\ [1000, 1000, 1000] 10) and bdd_above_ow: bdd_above.with_def (\(on _ with _ : \bdd'_above\ _)\ [1000, 1000, 1000] 10) and bdd_below_ow: bdd_below.with_def (\(on _ with _ : \bdd'_below\ _)\ [1000, 1000, 1000] 10) declare bdd.with[ud_with del] context preorder_ow begin abbreviation bdd_above :: "'ao set \ bool" where "bdd_above \ bdd_above_ow U (\\<^sub>o\<^sub>w)" abbreviation bdd_below :: "'ao set \ bool" where "bdd_below \ bdd_below_ow U (\\<^sub>o\<^sub>w)" end subsubsection\Transfer rules\ context includes lifting_syntax begin lemma partial_preordering_transfer[transfer_rule]: assumes [transfer_rule]: "right_total A" shows "((A ===> A ===> (=)) ===> (=)) (partial_preordering_ow (Collect (Domainp A))) partial_preordering" unfolding partial_preordering_ow_def partial_preordering_def apply transfer_prover_start apply transfer_step+ by blast lemma preordering_transfer[transfer_rule]: assumes [transfer_rule]: "right_total A" shows "((A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> (=)) (preordering_ow (Collect (Domainp A))) preordering" unfolding preordering_ow_def preordering_ow_axioms_def preordering_def preordering_axioms_def apply transfer_prover_start apply transfer_step+ by blast lemma preorder_transfer[transfer_rule]: assumes [transfer_rule]: "right_total A" shows "((A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> (=)) (preorder_ow (Collect (Domainp A))) class.preorder" unfolding preorder_ow_def class.preorder_def apply transfer_prover_start apply transfer_step+ by blast end subsubsection\Relativization\ context preordering_ow begin tts_context tts: (?'a to U) rewriting ctr_simps substituting preordering_ow_axioms eliminating through auto begin tts_lemma strict_implies_order: assumes "a \ U" and "b \ U" and "a \<^bold><\<^sub>o\<^sub>w b" shows "a \<^bold>\\<^sub>o\<^sub>w b" is preordering.strict_implies_order. tts_lemma irrefl: assumes "a \ U" shows "\a \<^bold><\<^sub>o\<^sub>w a" is preordering.irrefl. tts_lemma asym: assumes "a \ U" and "b \ U" and "a \<^bold><\<^sub>o\<^sub>w b" and "b \<^bold><\<^sub>o\<^sub>w a" shows False is preordering.asym. tts_lemma strict_trans1: assumes "a \ U" and "b \ U" and "c \ U" and "a \<^bold>\\<^sub>o\<^sub>w b" and "b \<^bold><\<^sub>o\<^sub>w c" shows "a \<^bold><\<^sub>o\<^sub>w c" is preordering.strict_trans1. tts_lemma strict_trans2: assumes "a \ U" and "b \ U" and "c \ U" and "a \<^bold><\<^sub>o\<^sub>w b" and "b \<^bold>\\<^sub>o\<^sub>w c" shows "a \<^bold><\<^sub>o\<^sub>w c" is preordering.strict_trans2. tts_lemma strict_trans: assumes "a \ U" and "b \ U" and "c \ U" and "a \<^bold><\<^sub>o\<^sub>w b" and "b \<^bold><\<^sub>o\<^sub>w c" shows "a \<^bold><\<^sub>o\<^sub>w c" is preordering.strict_trans. end end context preorder_ow begin tts_context tts: (?'a to U) rewriting ctr_simps substituting preorder_ow_axioms eliminating through auto begin tts_lemma less_irrefl: assumes "x \ U" shows "\ x <\<^sub>o\<^sub>w x" is preorder_class.less_irrefl. tts_lemma bdd_below_Ioc: assumes "a \ U" and "b \ U" shows "bdd_below {a<\<^sub>o\<^sub>w..b}" is preorder_class.bdd_below_Ioc. tts_lemma bdd_above_Ioc: assumes "a \ U" and "b \ U" shows "bdd_above {a<\<^sub>o\<^sub>w..b}" is preorder_class.bdd_above_Ioc. tts_lemma bdd_above_Iic: assumes "b \ U" shows "bdd_above {..\<^sub>o\<^sub>wb}" is preorder_class.bdd_above_Iic. tts_lemma bdd_above_Iio: assumes "b \ U" shows "bdd_above {..<\<^sub>o\<^sub>wb}" is preorder_class.bdd_above_Iio. tts_lemma bdd_below_Ici: assumes "a \ U" shows "bdd_below {a..\<^sub>o\<^sub>w}" is preorder_class.bdd_below_Ici. tts_lemma bdd_below_Ioi: assumes "a \ U" shows "bdd_below {a<\<^sub>o\<^sub>w..}" is preorder_class.bdd_below_Ioi. tts_lemma bdd_above_Icc: assumes "a \ U" and "b \ U" shows "bdd_above {a..\<^sub>o\<^sub>wb}" is preorder_class.bdd_above_Icc. tts_lemma bdd_above_Ioo: assumes "a \ U" and "b \ U" shows "bdd_above {a<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wb}" is preorder_class.bdd_above_Ioo. tts_lemma bdd_below_Icc: assumes "a \ U" and "b \ U" shows "bdd_below {a..\<^sub>o\<^sub>wb}" is preorder_class.bdd_below_Icc. tts_lemma bdd_below_Ioo: assumes "a \ U" and "b \ U" shows "bdd_below {a<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wb}" is preorder_class.bdd_below_Ioo. tts_lemma bdd_above_Ico: assumes "a \ U" and "b \ U" shows "bdd_above (on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {a.. U" and "b \ U" shows "bdd_below (on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {a.. U" shows "{a<\<^sub>o\<^sub>w..} \ {a..\<^sub>o\<^sub>w}" is preorder_class.Ioi_le_Ico. tts_lemma eq_refl: assumes "y \ U" and "x = y" shows "x \\<^sub>o\<^sub>w y" is preorder_class.eq_refl. tts_lemma less_imp_le: assumes "x \ U" and "y \ U" and "x <\<^sub>o\<^sub>w y" shows "x \\<^sub>o\<^sub>w y" is preorder_class.less_imp_le. tts_lemma less_not_sym: assumes "x \ U" and "y \ U" and "x <\<^sub>o\<^sub>w y" shows "\ y <\<^sub>o\<^sub>w x" is preorder_class.less_not_sym. tts_lemma less_imp_not_less: assumes "x \ U" and "y \ U" and "x <\<^sub>o\<^sub>w y" shows "(\ y <\<^sub>o\<^sub>w x) = True" is preorder_class.less_imp_not_less. tts_lemma less_asym': assumes "a \ U" and "b \ U" and "a <\<^sub>o\<^sub>w b" and "b <\<^sub>o\<^sub>w a" shows P is preorder_class.less_asym'. tts_lemma less_imp_triv: assumes "x \ U" and "y \ U" and "x <\<^sub>o\<^sub>w y" shows "(y <\<^sub>o\<^sub>w x \ P) = True" is preorder_class.less_imp_triv. tts_lemma less_trans: assumes "x \ U" and "y \ U" and "z \ U" and "x <\<^sub>o\<^sub>w y" and "y <\<^sub>o\<^sub>w z" shows "x <\<^sub>o\<^sub>w z" is preorder_class.less_trans. tts_lemma less_le_trans: assumes "x \ U" and "y \ U" and "z \ U" and "x <\<^sub>o\<^sub>w y" and "y \\<^sub>o\<^sub>w z" shows "x <\<^sub>o\<^sub>w z" is preorder_class.less_le_trans. tts_lemma le_less_trans: assumes "x \ U" and "y \ U" and "z \ U" and "x \\<^sub>o\<^sub>w y" and "y <\<^sub>o\<^sub>w z" shows "x <\<^sub>o\<^sub>w z" is preorder_class.le_less_trans. tts_lemma bdd_aboveI: assumes "A \ U" and "M \ U" and "\x. \x \ U; x \ A\ \ x \\<^sub>o\<^sub>w M" shows "bdd_above A" is preorder_class.bdd_aboveI. tts_lemma bdd_belowI: assumes "A \ U" and "m \ U" and "\x. \x \ U; x \ A\ \ m \\<^sub>o\<^sub>w x" shows "bdd_below A" is preorder_class.bdd_belowI. tts_lemma less_asym: assumes "x \ U" and "y \ U" and "x <\<^sub>o\<^sub>w y" and "\ P \ y <\<^sub>o\<^sub>w x" shows P is preorder_class.less_asym. -tts_lemma transp_less: "transp_on U (<\<^sub>o\<^sub>w)" - is preorder_class.transp_less. - -tts_lemma transp_le: "transp_on U (\\<^sub>o\<^sub>w)" - is preorder_class.transp_le. - -tts_lemma transp_gr: "transp_on U (\x y. y <\<^sub>o\<^sub>w x)" - is preorder_class.transp_gr. - -tts_lemma transp_ge: "transp_on U (\x y. y \\<^sub>o\<^sub>w x)" - is preorder_class.transp_ge. - tts_lemma bdd_above_Int1: assumes "A \ U" and "B \ U" and "bdd_above A" shows "bdd_above (A \ B)" is preorder_class.bdd_above_Int1. tts_lemma bdd_above_Int2: assumes "B \ U" and "A \ U" and "bdd_above B" shows "bdd_above (A \ B)" is preorder_class.bdd_above_Int2. tts_lemma bdd_below_Int1: assumes "A \ U" and "B \ U" and "bdd_below A" shows "bdd_below (A \ B)" is preorder_class.bdd_below_Int1. tts_lemma bdd_below_Int2: assumes "B \ U" and "A \ U" and "bdd_below B" shows "bdd_below (A \ B)" is preorder_class.bdd_below_Int2. tts_lemma bdd_above_mono: assumes "B \ U" and "bdd_above B" and "A \ B" shows "bdd_above A" is preorder_class.bdd_above_mono. tts_lemma bdd_below_mono: assumes "B \ U" and "bdd_below B" and "A \ B" shows "bdd_below A" is preorder_class.bdd_below_mono. tts_lemma atLeastAtMost_subseteq_atLeastLessThan_iff: assumes "a \ U" and "b \ U" and "c \ U" and "d \ U" shows "({a..\<^sub>o\<^sub>wb} \ (on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {c..\<^sub>o\<^sub>w b \ b <\<^sub>o\<^sub>w d \ c \\<^sub>o\<^sub>w a)" is preorder_class.atLeastAtMost_subseteq_atLeastLessThan_iff. tts_lemma atMost_subset_iff: assumes "x \ U" and "y \ U" shows "({..\<^sub>o\<^sub>wx} \ {..\<^sub>o\<^sub>wy}) = (x \\<^sub>o\<^sub>w y)" is Set_Interval.atMost_subset_iff. tts_lemma single_Diff_lessThan: assumes "k \ U" shows "{k} - {..<\<^sub>o\<^sub>wk} = {k}" is Set_Interval.single_Diff_lessThan. tts_lemma atLeast_subset_iff: assumes "x \ U" and "y \ U" shows "({x..\<^sub>o\<^sub>w} \ {y..\<^sub>o\<^sub>w}) = (y \\<^sub>o\<^sub>w x)" is Set_Interval.atLeast_subset_iff. tts_lemma atLeastatMost_psubset_iff: assumes "a \ U" and "b \ U" and "c \ U" and "d \ U" shows "({a..\<^sub>o\<^sub>wb} \ {c..\<^sub>o\<^sub>wd}) = (c \\<^sub>o\<^sub>w d \ (\ a \\<^sub>o\<^sub>w b \ c \\<^sub>o\<^sub>w a \ b \\<^sub>o\<^sub>w d \ (c <\<^sub>o\<^sub>w a \ b <\<^sub>o\<^sub>w d)))" is preorder_class.atLeastatMost_psubset_iff. tts_lemma not_empty_eq_Iic_eq_empty: assumes "h \ U" shows "{} \ {..\<^sub>o\<^sub>wh}" is preorder_class.not_empty_eq_Iic_eq_empty. tts_lemma atLeastatMost_subset_iff: assumes "a \ U" and "b \ U" and "c \ U" and "d \ U" shows "({a..\<^sub>o\<^sub>wb} \ {c..\<^sub>o\<^sub>wd}) = (\ a \\<^sub>o\<^sub>w b \ b \\<^sub>o\<^sub>w d \ c \\<^sub>o\<^sub>w a)" is preorder_class.atLeastatMost_subset_iff. tts_lemma Icc_subset_Ici_iff: assumes "l \ U" and "h \ U" and "l' \ U" shows "({l..\<^sub>o\<^sub>wh} \ {l'..\<^sub>o\<^sub>w}) = (\ l \\<^sub>o\<^sub>w h \ l' \\<^sub>o\<^sub>w l)" is preorder_class.Icc_subset_Ici_iff. tts_lemma Icc_subset_Iic_iff: assumes "l \ U" and "h \ U" and "h' \ U" shows "({l..\<^sub>o\<^sub>wh} \ {..\<^sub>o\<^sub>wh'}) = (\ l \\<^sub>o\<^sub>w h \ h \\<^sub>o\<^sub>w h')" is preorder_class.Icc_subset_Iic_iff. tts_lemma not_empty_eq_Ici_eq_empty: assumes "l \ U" shows "{} \ {l..\<^sub>o\<^sub>w}" is preorder_class.not_empty_eq_Ici_eq_empty. tts_lemma not_Ici_eq_empty: assumes "l \ U" shows "{l..\<^sub>o\<^sub>w} \ {}" is preorder_class.not_Ici_eq_empty. tts_lemma not_Iic_eq_empty: assumes "h \ U" shows "{..\<^sub>o\<^sub>wh} \ {}" is preorder_class.not_Iic_eq_empty. tts_lemma atLeastatMost_empty_iff2: assumes "a \ U" and "b \ U" shows "({} = {a..\<^sub>o\<^sub>wb}) = (\ a \\<^sub>o\<^sub>w b)" is preorder_class.atLeastatMost_empty_iff2. tts_lemma atLeastLessThan_empty_iff2: assumes "a \ U" and "b \ U" shows "({} = (on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {a.. a <\<^sub>o\<^sub>w b)" is preorder_class.atLeastLessThan_empty_iff2. tts_lemma greaterThanAtMost_empty_iff2: assumes "k \ U" and "l \ U" shows "({} = {k<\<^sub>o\<^sub>w..l}) = (\ k <\<^sub>o\<^sub>w l)" is preorder_class.greaterThanAtMost_empty_iff2. tts_lemma atLeastatMost_empty_iff: assumes "a \ U" and "b \ U" shows "({a..\<^sub>o\<^sub>wb} = {}) = (\ a \\<^sub>o\<^sub>w b)" is preorder_class.atLeastatMost_empty_iff. tts_lemma atLeastLessThan_empty_iff: assumes "a \ U" and "b \ U" shows "((on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {a.. a <\<^sub>o\<^sub>w b)" is preorder_class.atLeastLessThan_empty_iff. tts_lemma greaterThanAtMost_empty_iff: assumes "k \ U" and "l \ U" shows "({k<\<^sub>o\<^sub>w..l} = {}) = (\ k <\<^sub>o\<^sub>w l)" is preorder_class.greaterThanAtMost_empty_iff. end tts_context tts: (?'a to U) substituting preorder_ow_axioms begin tts_lemma bdd_above_empty: assumes "U \ {}" shows "bdd_above {}" is preorder_class.bdd_above_empty. tts_lemma bdd_below_empty: assumes "U \ {}" shows "bdd_below {}" is preorder_class.bdd_below_empty. end tts_context tts: (?'a to U) and (?'b to \U\<^sub>2::'a set\) rewriting ctr_simps substituting preorder_ow_axioms eliminating through (auto intro: bdd_above_empty bdd_below_empty) begin tts_lemma bdd_belowI2: assumes "A \ U\<^sub>2" and "m \ U" and "\x\U\<^sub>2. f x \ U" and "\x. x \ A \ m \\<^sub>o\<^sub>w f x" shows "bdd_below (f ` A)" given preorder_class.bdd_belowI2 by blast tts_lemma bdd_aboveI2: assumes "A \ U\<^sub>2" and "\x\U\<^sub>2. f x \ U" and "M \ U" and "\x. x \ A \ f x \\<^sub>o\<^sub>w M" shows "bdd_above (f ` A)" given preorder_class.bdd_aboveI2 by blast end end subsection\Partial orders\ subsubsection\Definitions and common properties\ locale ordering_ow = partial_preordering_ow U le for U :: "'ao set" and le :: "'ao \ 'ao \ bool" (infix "\<^bold>\\<^sub>o\<^sub>w" 50) + fixes ls :: "'ao \ 'ao \ bool" (infix "\<^bold><\<^sub>o\<^sub>w" 50) assumes strict_iff_order: "\ a \ U; b \ U \ \ a \<^bold><\<^sub>o\<^sub>w b \ a \<^bold>\\<^sub>o\<^sub>w b \ a \ b" and antisym: "\ a \ U; b \ U; a \<^bold>\\<^sub>o\<^sub>w b; b \<^bold>\\<^sub>o\<^sub>w a \ \ a = b" begin notation le (infix "\<^bold>\\<^sub>o\<^sub>w" 50) and ls (infix "\<^bold><\<^sub>o\<^sub>w" 50) sublocale preordering_ow U \(\<^bold>\\<^sub>o\<^sub>w)\ \(\<^bold><\<^sub>o\<^sub>w)\ using local.antisym strict_iff_order by unfold_locales blast end locale order_ow = preorder_ow U le ls for U :: "'ao set" and le ls + assumes antisym: "\ x \ U; y \ U; x \\<^sub>o\<^sub>w y; y \\<^sub>o\<^sub>w x \ \ x = y" begin sublocale order: ordering_ow U \(\\<^sub>o\<^sub>w)\ \(<\<^sub>o\<^sub>w)\ + dual_order: ordering_ow U \(\\<^sub>o\<^sub>w)\ \(>\<^sub>o\<^sub>w)\ apply unfold_locales subgoal by (force simp: less_le_not_le antisym) subgoal by (simp add: antisym) subgoal by (force simp: less_le_not_le antisym) subgoal by (simp add: antisym) done no_notation le (infix "\<^bold>\\<^sub>o\<^sub>w" 50) and ls (infix "\<^bold><\<^sub>o\<^sub>w" 50) end locale ord_order_ow = ord\<^sub>1: ord_ow U\<^sub>1 le\<^sub>1 ls\<^sub>1 + ord\<^sub>2: order_ow U\<^sub>2 le\<^sub>2 ls\<^sub>2 for U\<^sub>1 :: "'ao set" and le\<^sub>1 ls\<^sub>1 and U\<^sub>2 :: "'bo set" and le\<^sub>2 ls\<^sub>2 sublocale ord_order_ow \ ord_preorder_ow .. locale preorder_order_ow = ord\<^sub>1: preorder_ow U\<^sub>1 le\<^sub>1 ls\<^sub>1 + ord\<^sub>2: order_ow U\<^sub>2 le\<^sub>2 ls\<^sub>2 for U\<^sub>1 :: "'ao set" and le\<^sub>1 ls\<^sub>1 and U\<^sub>2 :: "'bo set" and le\<^sub>2 ls\<^sub>2 sublocale preorder_order_ow \ preorder_pair_ow .. locale order_pair_ow = ord\<^sub>1: order_ow U\<^sub>1 le\<^sub>1 ls\<^sub>1 + ord\<^sub>2: order_ow U\<^sub>2 le\<^sub>2 ls\<^sub>2 for U\<^sub>1 :: "'ao set" and le\<^sub>1 ls\<^sub>1 and U\<^sub>2 :: "'bo set" and le\<^sub>2 ls\<^sub>2 sublocale order_pair_ow \ preorder_order_ow .. ud \monoseq\ (\(with _ : \monoseq\ _)\ [1000, 1000] 10) ctr relativization synthesis ctr_simps assumes [transfer_domain_rule, transfer_rule]: "Domainp (B::'c\'d\bool) = (\x. x \ U\<^sub>2)" and [transfer_rule]: "right_total B" trp (?'b \A::'a\'b\bool\) and (?'a B) in monoseq_ow: monoseq.with_def subsubsection\Transfer rules\ context includes lifting_syntax begin lemma ordering_transfer[transfer_rule]: assumes [transfer_rule]: "bi_unique A" "right_total A" shows "((A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> (=)) (ordering_ow (Collect (Domainp A))) ordering" unfolding ordering_ow_def ordering_ow_axioms_def ordering_def ordering_axioms_def apply transfer_prover_start apply transfer_step+ unfolding Ball_Collect[symmetric] by (intro ext HOL.arg_cong2[where f="(\)"]) auto lemma order_transfer[transfer_rule]: assumes [transfer_rule]: "bi_unique A" "right_total A" shows "((A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> (=)) (order_ow (Collect (Domainp A))) class.order" unfolding order_ow_def class.order_def order_ow_axioms_def class.order_axioms_def apply transfer_prover_start apply transfer_step+ by simp end subsubsection\Relativization\ context ordering_ow begin tts_context tts: (?'a to U) rewriting ctr_simps substituting ordering_ow_axioms eliminating through simp begin tts_lemma strict_implies_not_eq: assumes "a \ U" and "b \ U" and "a \<^bold><\<^sub>o\<^sub>w b" shows "a \ b" is ordering.strict_implies_not_eq. tts_lemma order_iff_strict: assumes "a \ U" and "b \ U" shows "(a \<^bold>\\<^sub>o\<^sub>w b) = (a \<^bold><\<^sub>o\<^sub>w b \ a = b)" is ordering.order_iff_strict. tts_lemma not_eq_order_implies_strict: assumes "a \ U" and "b \ U" and "a \ b" and "a \<^bold>\\<^sub>o\<^sub>w b" shows "a \<^bold><\<^sub>o\<^sub>w b" is ordering.not_eq_order_implies_strict. tts_lemma eq_iff: assumes "a \ U" and "b \ U" shows "(a = b) = (a \<^bold>\\<^sub>o\<^sub>w b \ b \<^bold>\\<^sub>o\<^sub>w a)" is ordering.eq_iff. end end context order_ow begin tts_context tts: (?'a to U) rewriting ctr_simps substituting order_ow_axioms eliminating through clarsimp begin tts_lemma atLeastAtMost_singleton: assumes "a \ U" shows "{a..\<^sub>o\<^sub>wa} = {a}" is order_class.atLeastAtMost_singleton. tts_lemma less_imp_neq: assumes "y \ U" and "x <\<^sub>o\<^sub>w y" shows "x \ y" is order_class.less_imp_neq. tts_lemma atLeastatMost_empty: assumes "b \ U" and "a \ U" and "b <\<^sub>o\<^sub>w a" shows "{a..\<^sub>o\<^sub>wb} = {}" is order_class.atLeastatMost_empty. tts_lemma less_imp_not_eq: assumes "y \ U" and "x <\<^sub>o\<^sub>w y" shows "(x = y) = False" is order_class.less_imp_not_eq. tts_lemma less_imp_not_eq2: assumes "y \ U" and "x <\<^sub>o\<^sub>w y" shows "(y = x) = False" is order_class.less_imp_not_eq2. tts_lemma atLeastLessThan_empty: assumes "b \ U" and "a \ U" and "b \\<^sub>o\<^sub>w a" shows "(on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {a.. U" and "k \ U" and "l \\<^sub>o\<^sub>w k" shows "{k<\<^sub>o\<^sub>w..l} = {}" is order_class.greaterThanAtMost_empty. tts_lemma antisym_conv1: assumes "x \ U" and "y \ U" and "\ x <\<^sub>o\<^sub>w y" shows "(x \\<^sub>o\<^sub>w y) = (x = y)" is order_class.antisym_conv1. tts_lemma antisym_conv2: assumes "x \ U" and "y \ U" and "x \\<^sub>o\<^sub>w y" shows "(\ x <\<^sub>o\<^sub>w y) = (x = y)" is order_class.antisym_conv2. tts_lemma greaterThanLessThan_empty: assumes "l \ U" and "k \ U" and "l \\<^sub>o\<^sub>w k" shows "{k<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wl} = {}" is order_class.greaterThanLessThan_empty. tts_lemma atLeastLessThan_eq_atLeastAtMost_diff: assumes "a \ U" and "b \ U" shows "(on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {a..o\<^sub>wb} - {b}" is order_class.atLeastLessThan_eq_atLeastAtMost_diff. tts_lemma greaterThanAtMost_eq_atLeastAtMost_diff: assumes "a \ U" and "b \ U" shows "{a<\<^sub>o\<^sub>w..b} = {a..\<^sub>o\<^sub>wb} - {a}" is order_class.greaterThanAtMost_eq_atLeastAtMost_diff. tts_lemma less_separate: assumes "x \ U" and "y \ U" and "x <\<^sub>o\<^sub>w y" shows "\x'\U. \y'\U. x \ {..<\<^sub>o\<^sub>wx'} \ y \ {y'<\<^sub>o\<^sub>w..} \ {..<\<^sub>o\<^sub>wx'} \ {y'<\<^sub>o\<^sub>w..} = {}" is order_class.less_separate. tts_lemma order_iff_strict: assumes "a \ U" and "b \ U" shows "(a \\<^sub>o\<^sub>w b) = (a <\<^sub>o\<^sub>w b \ a = b)" is order_class.order.order_iff_strict. tts_lemma le_less: assumes "x \ U" and "y \ U" shows "(x \\<^sub>o\<^sub>w y) = (x <\<^sub>o\<^sub>w y \ x = y)" is order_class.le_less. tts_lemma strict_iff_order: assumes "a \ U" and "b \ U" shows "(a <\<^sub>o\<^sub>w b) = (a \\<^sub>o\<^sub>w b \ a \ b)" is order_class.order.strict_iff_order. tts_lemma less_le: assumes "x \ U" and "y \ U" shows "(x <\<^sub>o\<^sub>w y) = (x \\<^sub>o\<^sub>w y \ x \ y)" is order_class.less_le. tts_lemma atLeastAtMost_singleton': assumes "b \ U" and "a = b" shows "{a..\<^sub>o\<^sub>wb} = {a}" is order_class.atLeastAtMost_singleton'. tts_lemma le_imp_less_or_eq: assumes "x \ U" and "y \ U" and "x \\<^sub>o\<^sub>w y" shows "x <\<^sub>o\<^sub>w y \ x = y" is order_class.le_imp_less_or_eq. tts_lemma antisym_conv: assumes "y \ U" and "x \ U" and "y \\<^sub>o\<^sub>w x" shows "(x \\<^sub>o\<^sub>w y) = (x = y)" is order_class.antisym_conv. tts_lemma le_neq_trans: assumes "a \ U" and "b \ U" and "a \\<^sub>o\<^sub>w b" and "a \ b" shows "a <\<^sub>o\<^sub>w b" is order_class.le_neq_trans. tts_lemma neq_le_trans: assumes "a \ U" and "b \ U" and "a \ b" and "a \\<^sub>o\<^sub>w b" shows "a <\<^sub>o\<^sub>w b" is order_class.neq_le_trans. tts_lemma Iio_Int_singleton: assumes "k \ U" and "x \ U" shows "{..<\<^sub>o\<^sub>wk} \ {x} = (if x <\<^sub>o\<^sub>w k then {x} else {})" is order_class.Iio_Int_singleton. tts_lemma atLeastAtMost_singleton_iff: assumes "a \ U" and "b \ U" and "c \ U" shows "({a..\<^sub>o\<^sub>wb} = {c}) = (a = b \ b = c)" is order_class.atLeastAtMost_singleton_iff. tts_lemma Icc_eq_Icc: assumes "l \ U" and "h \ U" and "l' \ U" and "h' \ U" shows "({l..\<^sub>o\<^sub>wh} = {l'..\<^sub>o\<^sub>wh'}) = (h = h' \ l = l' \ \ l' \\<^sub>o\<^sub>w h' \ \ l \\<^sub>o\<^sub>w h)" is order_class.Icc_eq_Icc. tts_lemma lift_Suc_mono_less_iff: assumes "range f \ U" and "\n. f n <\<^sub>o\<^sub>w f (Suc n)" shows "(f n <\<^sub>o\<^sub>w f m) = (n < m)" is order_class.lift_Suc_mono_less_iff. tts_lemma lift_Suc_mono_less: assumes "range f \ U" and "\n. f n <\<^sub>o\<^sub>w f (Suc n)" and "n < n'" shows "f n <\<^sub>o\<^sub>w f n'" is order_class.lift_Suc_mono_less. tts_lemma lift_Suc_mono_le: assumes "range f \ U" and "\n. f n \\<^sub>o\<^sub>w f (Suc n)" and "n \ n'" shows "f n \\<^sub>o\<^sub>w f n'" is order_class.lift_Suc_mono_le. tts_lemma lift_Suc_antimono_le: assumes "range f \ U" and "\n. f (Suc n) \\<^sub>o\<^sub>w f n" and "n \ n'" shows "f n' \\<^sub>o\<^sub>w f n" is order_class.lift_Suc_antimono_le. tts_lemma ivl_disj_int_two: assumes "l \ U" and "m \ U" and "u \ U" shows "{l<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wm} \ (on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {m..o\<^sub>w..m} \ {m<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wu} = {}" "(on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {l.. (on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {m..o\<^sub>wm} \ {m<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wu} = {}" "{l<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wm} \ {m..\<^sub>o\<^sub>wu} = {}" "{l<\<^sub>o\<^sub>w..m} \ {m<\<^sub>o\<^sub>w..u} = {}" "(on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {l.. {m..\<^sub>o\<^sub>wu} = {}" "{l..\<^sub>o\<^sub>wm} \ {m<\<^sub>o\<^sub>w..u} = {}" is Set_Interval.ivl_disj_int_two. tts_lemma ivl_disj_int_one: assumes "l \ U" and "u \ U" shows "{..\<^sub>o\<^sub>wl} \ {l<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wu} = {}" "{..<\<^sub>o\<^sub>wl} \ (on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {l..o\<^sub>wl} \ {l<\<^sub>o\<^sub>w..u} = {}" "{..<\<^sub>o\<^sub>wl} \ {l..\<^sub>o\<^sub>wu} = {}" "{l<\<^sub>o\<^sub>w..u} \ {u<\<^sub>o\<^sub>w..} = {}" "{l<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wu} \ {u..\<^sub>o\<^sub>w} = {}" "{l..\<^sub>o\<^sub>wu} \ {u<\<^sub>o\<^sub>w..} = {}" "(on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {l.. {u..\<^sub>o\<^sub>w} = {}" is Set_Interval.ivl_disj_int_one. tts_lemma min_absorb2: assumes "y \ U" and "x \ U" and "y \\<^sub>o\<^sub>w x" shows "local.min x y = y" is Orderings.min_absorb2. tts_lemma max_absorb1: assumes "y \ U" and "x \ U" and "y \\<^sub>o\<^sub>w x" shows "local.max x y = x" is Orderings.max_absorb1. tts_lemma atMost_Int_atLeast: assumes "n \ U" shows "{..\<^sub>o\<^sub>wn} \ {n..\<^sub>o\<^sub>w} = {n}" is Set_Interval.atMost_Int_atLeast. tts_lemma monoseq_Suc: assumes "range X \ U" shows "(with (\\<^sub>o\<^sub>w) : \monoseq\ X) = ((\x. X x \\<^sub>o\<^sub>w X (Suc x)) \ (\x. X (Suc x) \\<^sub>o\<^sub>w X x))" is Topological_Spaces.monoseq_Suc. tts_lemma mono_SucI2: assumes "range X \ U" and "\x. X (Suc x) \\<^sub>o\<^sub>w X x" shows "with (\\<^sub>o\<^sub>w) : \monoseq\ X" is Topological_Spaces.mono_SucI2. tts_lemma mono_SucI1: assumes "range X \ U" and "\x. X x \\<^sub>o\<^sub>w X (Suc x)" shows "with (\\<^sub>o\<^sub>w) : \monoseq\ X" is Topological_Spaces.mono_SucI1. tts_lemma monoI2: assumes "range X \ U" and "\x y. x \ y \ X y \\<^sub>o\<^sub>w X x" shows "with (\\<^sub>o\<^sub>w) : \monoseq\ X" is Topological_Spaces.monoI2. tts_lemma monoI1: assumes "range X \ U" and "\x y. x \ y \ X x \\<^sub>o\<^sub>w X y" shows "with (\\<^sub>o\<^sub>w) : \monoseq\ X" is Topological_Spaces.monoI1. end tts_context tts: (?'a to U) rewriting ctr_simps substituting order_ow_axioms eliminating through clarsimp begin tts_lemma ex_min_if_finite: assumes "S \ U" and "finite S" and "S \ {}" shows "\x\S. \ (\y\S. y <\<^sub>o\<^sub>w x)" is Lattices_Big.ex_min_if_finite. end tts_context tts: (?'a to U) sbterms: (\(\)::['a::order, 'a::order] \ bool\ to \(\\<^sub>o\<^sub>w)\) and (\(<)::['a::order, 'a::order] \ bool\ to \(<\<^sub>o\<^sub>w)\) substituting order_ow_axioms eliminating through clarsimp begin tts_lemma xt1: shows "\a = b; c <\<^sub>o\<^sub>w b\ \ c <\<^sub>o\<^sub>w a" "\b <\<^sub>o\<^sub>w a; b = c\ \ c <\<^sub>o\<^sub>w a" "\a = b; c \\<^sub>o\<^sub>w b\ \ c \\<^sub>o\<^sub>w a" "\b \\<^sub>o\<^sub>w a; b = c\ \ c \\<^sub>o\<^sub>w a" "\y \ U; x \ U; y \\<^sub>o\<^sub>w x; x \\<^sub>o\<^sub>w y\ \ x = y" "\y \ U; x \ U; z \ U; y \\<^sub>o\<^sub>w x; z \\<^sub>o\<^sub>w y\ \ z \\<^sub>o\<^sub>w x" "\y \ U; x \ U; z \ U; y <\<^sub>o\<^sub>w x; z \\<^sub>o\<^sub>w y\ \ z <\<^sub>o\<^sub>w x" "\y \ U; x \ U; z \ U; y \\<^sub>o\<^sub>w x; z <\<^sub>o\<^sub>w y\ \ z <\<^sub>o\<^sub>w x" "\b \ U; a \ U; b <\<^sub>o\<^sub>w a; a <\<^sub>o\<^sub>w b\ \ P" "\y \ U; x \ U; z \ U; y <\<^sub>o\<^sub>w x; z <\<^sub>o\<^sub>w y\ \ z <\<^sub>o\<^sub>w x" "\b \ U; a \ U; b \\<^sub>o\<^sub>w a; a \ b\ \ b <\<^sub>o\<^sub>w a" "\a \ U; b \ U; a \ b; b \\<^sub>o\<^sub>w a\ \ b <\<^sub>o\<^sub>w a" "\ b \ U; c \ U; a = f b; c <\<^sub>o\<^sub>w b; \x y. \x \ U; y \ U; y <\<^sub>o\<^sub>w x\ \ f y <\<^sub>o\<^sub>w f x \ \ f c <\<^sub>o\<^sub>w a" "\ b \ U; a \ U; b <\<^sub>o\<^sub>w a; f b = c; \x y. \x \ U; y \ U; y <\<^sub>o\<^sub>w x\ \ f y <\<^sub>o\<^sub>w f x \ \ c <\<^sub>o\<^sub>w f a" "\ b \ U; c \ U; a = f b; c \\<^sub>o\<^sub>w b; \x y. \x \ U; y \ U; y \\<^sub>o\<^sub>w x\ \ f y \\<^sub>o\<^sub>w f x \ \ f c \\<^sub>o\<^sub>w a" "\ b \ U; a \ U; b \\<^sub>o\<^sub>w a; f b = c; \x y. \x \ U; y \ U; y \\<^sub>o\<^sub>w x\ \ f y \\<^sub>o\<^sub>w f x \ \ c \\<^sub>o\<^sub>w f a" is Orderings.xt1. end end context ord_order_ow begin tts_context tts: (?'a to U\<^sub>2) and (?'b to U\<^sub>1) sbterms: (\(\)::[?'a::order, ?'a::order] \ bool\ to \(\\<^sub>o\<^sub>w\<^sub>.\<^sub>2)\) and (\(<)::[?'a::order, ?'a::order] \ bool\ to \(<\<^sub>o\<^sub>w\<^sub>.\<^sub>2)\) and (\(\)::[?'b::ord, ?'b::ord] \ bool\ to \(\\<^sub>o\<^sub>w\<^sub>.\<^sub>1)\) and (\(<)::[?'b::ord, ?'b::ord] \ bool\ to \(<\<^sub>o\<^sub>w\<^sub>.\<^sub>1)\) rewriting ctr_simps substituting ord\<^sub>2.order_ow_axioms eliminating through clarsimp begin tts_lemma xt2: assumes "\x\U\<^sub>1. f x \ U\<^sub>2" and "b \ U\<^sub>1" and "a \ U\<^sub>2" and "c \ U\<^sub>1" and "f b \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 a" and "c \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 b" and "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; y \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 x\ \ f y \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f x" shows "f c \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 a" is Orderings.xt2. tts_lemma xt6: assumes "\x\U\<^sub>1. f x \ U\<^sub>2" and "b \ U\<^sub>1" and "a \ U\<^sub>2" and "c \ U\<^sub>1" and "f b \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 a" and "c <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 b" and "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; y <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 x\ \ f y <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f x" shows "f c <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 a" is Orderings.xt6. end end context order_pair_ow begin tts_context tts: (?'a to U\<^sub>1) and (?'b to U\<^sub>2) sbterms: (\(\)::[?'a::order, ?'a::order] \ bool\ to \(\\<^sub>o\<^sub>w\<^sub>.\<^sub>1)\) and (\(<)::[?'a::order, ?'a::order] \ bool\ to \(<\<^sub>o\<^sub>w\<^sub>.\<^sub>1)\) and (\(\)::[?'b::order, ?'b::order] \ bool\ to \(\\<^sub>o\<^sub>w\<^sub>.\<^sub>2)\) and (\(<)::[?'b::order, ?'b::order] \ bool\ to \(<\<^sub>o\<^sub>w\<^sub>.\<^sub>2)\) rewriting ctr_simps substituting ord\<^sub>1.order_ow_axioms and ord\<^sub>2.order_ow_axioms eliminating through clarsimp begin tts_lemma xt3: assumes "b \ U\<^sub>1" and "a \ U\<^sub>1" and "c \ U\<^sub>2" and "\x\U\<^sub>1. f x \ U\<^sub>2" and "b \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 a" and "c \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f b" and "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; y \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 x\ \ f y \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f x" shows "c \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f a" is Orderings.xt3. tts_lemma xt4: assumes "\x\U\<^sub>2. f x \ U\<^sub>1" and "b \ U\<^sub>2" and "a \ U\<^sub>1" and "c \ U\<^sub>2" and "f b <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 a" and "c \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 b" and "\x y. \x \ U\<^sub>2; y \ U\<^sub>2; y \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 x\ \ f y \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 f x" shows "f c <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 a" is Orderings.xt4. tts_lemma xt5: assumes "b \ U\<^sub>1" and "a \ U\<^sub>1" and "c \ U\<^sub>2" and "\x\U\<^sub>1. f x \ U\<^sub>2" and "b <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 a" and "c \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f b" and "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; y <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 x\ \ f y <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f x" shows "c <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f a" is Orderings.xt5. tts_lemma xt7: assumes "b \ U\<^sub>1" and "a \ U\<^sub>1" and "c \ U\<^sub>2" and "\x\U\<^sub>1. f x \ U\<^sub>2" and "b \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 a" and "c <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f b" and "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; y \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 x\ \ f y \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f x" shows "c <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f a" is Orderings.xt7. tts_lemma xt8: assumes "\x\U\<^sub>2. f x \ U\<^sub>1" and "b \ U\<^sub>2" and "a \ U\<^sub>1" and "c \ U\<^sub>2" and "f b <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 a" and "c <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 b" and "\x y. \x \ U\<^sub>2; y \ U\<^sub>2; y <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 x\ \ f y <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 f x" shows "f c <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 a" is Orderings.xt8. tts_lemma xt9: assumes "b \ U\<^sub>1" and "a \ U\<^sub>1" and "c \ U\<^sub>2" and "\x\U\<^sub>1. f x \ U\<^sub>2" and "b <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 a" and "c <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f b" and "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; y <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 x\ \ f y <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f x" shows "c <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f a" is Orderings.xt9. end tts_context tts: (?'a to U\<^sub>1) and (?'b to U\<^sub>2) sbterms: (\(\)::[?'a::order, ?'a::order] \ bool\ to \(\\<^sub>o\<^sub>w\<^sub>.\<^sub>1)\) and (\(<)::[?'a::order, ?'a::order] \ bool\ to \(<\<^sub>o\<^sub>w\<^sub>.\<^sub>1)\) and (\(\)::[?'b::order, ?'b::order] \ bool\ to \(\\<^sub>o\<^sub>w\<^sub>.\<^sub>2)\) and (\(<)::[?'b::order, ?'b::order] \ bool\ to \(<\<^sub>o\<^sub>w\<^sub>.\<^sub>2)\) rewriting ctr_simps substituting ord\<^sub>1.order_ow_axioms and ord\<^sub>2.order_ow_axioms eliminating through simp begin tts_lemma order_less_subst1: assumes "a \ U\<^sub>1" and "\x\U\<^sub>2. f x \ U\<^sub>1" and "b \ U\<^sub>2" and "c \ U\<^sub>2" and "a <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 f b" and "b <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 c" and "\x y. \x \ U\<^sub>2; y \ U\<^sub>2; x <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 y\ \ f x <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 f y" shows "a <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 f c" is Orderings.order_less_subst1. tts_lemma order_subst1: assumes "a \ U\<^sub>1" and "\x\U\<^sub>2. f x \ U\<^sub>1" and "b \ U\<^sub>2" and "c \ U\<^sub>2" and "a \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 f b" and "b \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 c" and "\x y. \x \ U\<^sub>2; y \ U\<^sub>2; x \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 y\ \ f x \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 f y" shows "a \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 f c" is Orderings.order_subst1. end tts_context tts: (?'a to U\<^sub>1) and (?'c to U\<^sub>2) sbterms: (\(\)::[?'a::order, ?'a::order] \ bool\ to \(\\<^sub>o\<^sub>w\<^sub>.\<^sub>1)\) and (\(<)::[?'a::order, ?'a::order] \ bool\ to \(<\<^sub>o\<^sub>w\<^sub>.\<^sub>1)\) and (\(\)::[?'c::order, ?'c::order] \ bool\ to \(\\<^sub>o\<^sub>w\<^sub>.\<^sub>2)\) and (\(<)::[?'c::order, ?'c::order] \ bool\ to \(<\<^sub>o\<^sub>w\<^sub>.\<^sub>2)\) rewriting ctr_simps substituting ord\<^sub>1.order_ow_axioms and ord\<^sub>2.order_ow_axioms eliminating through simp begin tts_lemma order_less_le_subst2: assumes "a \ U\<^sub>1" and "b \ U\<^sub>1" and "\x\U\<^sub>1. f x \ U\<^sub>2" and "c \ U\<^sub>2" and "a <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 b" and "f b \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 c" and "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; x <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 y\ \ f x <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f y" shows "f a <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 c" is Orderings.order_less_le_subst2. tts_lemma order_le_less_subst2: assumes "a \ U\<^sub>1" and "b \ U\<^sub>1" and "\x\U\<^sub>1. f x \ U\<^sub>2" and "c \ U\<^sub>2" and "a \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 b" and "f b <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 c" and "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; x \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 y\ \ f x \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f y" shows "f a <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 c" is Orderings.order_le_less_subst2. tts_lemma order_less_subst2: assumes "a \ U\<^sub>1" and "b \ U\<^sub>1" and "\x\U\<^sub>1. f x \ U\<^sub>2" and "c \ U\<^sub>2" and "a <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 b" and "f b <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 c" and "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; x <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 y\ \ f x <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f y" shows "f a <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 c" is Orderings.order_less_subst2. tts_lemma order_subst2: assumes "a \ U\<^sub>1" and "b \ U\<^sub>1" and "\x\U\<^sub>1. f x \ U\<^sub>2" and "c \ U\<^sub>2" and "a \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 b" and "f b \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 c" and "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; x \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 y\ \ f x \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f y" shows "f a \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 c" is Orderings.order_subst2. end end subsection\Dense orders\ subsubsection\Definitions and common properties\ locale dense_order_ow = order_ow U le ls for U :: "'ao set" and le ls + assumes dense: "\ x \ U; y \ U; x <\<^sub>o\<^sub>w y \ \ (\z\U. x <\<^sub>o\<^sub>w z \ z <\<^sub>o\<^sub>w y)" subsubsection\Transfer rules\ context includes lifting_syntax begin lemma dense_order_transfer[transfer_rule]: assumes [transfer_rule]: "bi_unique A" "right_total A" shows "((A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> (=)) (dense_order_ow (Collect (Domainp A))) class.dense_order" unfolding dense_order_ow_def class.dense_order_def dense_order_ow_axioms_def class.dense_order_axioms_def apply transfer_prover_start apply transfer_step+ by simp end subsection\ Partial orders with the greatest element and partial orders with the least elements \ subsubsection\Definitions and common properties\ locale ordering_top_ow = ordering_ow U le ls for U :: "'ao set" and le ls + fixes top :: "'ao" ("\<^bold>\\<^sub>o\<^sub>w") assumes top_closed[simp]: "\<^bold>\\<^sub>o\<^sub>w \ U" assumes extremum[simp]: "a \ U \ a \<^bold>\\<^sub>o\<^sub>w \<^bold>\\<^sub>o\<^sub>w" begin notation top ("\<^bold>\\<^sub>o\<^sub>w") end locale bot_ow = fixes U :: "'ao set" and bot (\\\<^sub>o\<^sub>w\) assumes bot_closed[simp]: "\\<^sub>o\<^sub>w \ U" begin notation bot (\\\<^sub>o\<^sub>w\) end locale bot_pair_ow = ord\<^sub>1: bot_ow U\<^sub>1 bot\<^sub>1 + ord\<^sub>2: bot_ow U\<^sub>2 bot\<^sub>2 for U\<^sub>1 :: "'ao set" and bot\<^sub>1 and U\<^sub>2 :: "'bo set" and bot\<^sub>2 begin notation bot\<^sub>1 (\\\<^sub>o\<^sub>w\<^sub>.\<^sub>1\) notation bot\<^sub>2 (\\\<^sub>o\<^sub>w\<^sub>.\<^sub>2\) end locale order_bot_ow = order_ow U le ls + bot_ow U bot for U :: "'ao set" and bot le ls + assumes bot_least: "a \ U \ \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w a" begin sublocale bot: ordering_top_ow U \(\\<^sub>o\<^sub>w)\ \(>\<^sub>o\<^sub>w)\ \\\<^sub>o\<^sub>w\ apply unfold_locales subgoal by simp subgoal by (simp add: bot_least) done no_notation le (infix "\<^bold>\\<^sub>o\<^sub>w" 50) and ls (infix "\<^bold><\<^sub>o\<^sub>w" 50) and top ("\<^bold>\\<^sub>o\<^sub>w") end locale order_bot_pair_ow = ord\<^sub>1: order_bot_ow U\<^sub>1 bot\<^sub>1 le\<^sub>1 ls\<^sub>1 + ord\<^sub>2: order_bot_ow U\<^sub>2 bot\<^sub>2 le\<^sub>2 ls\<^sub>2 for U\<^sub>1 :: "'ao set" and bot\<^sub>1 le\<^sub>1 ls\<^sub>1 and U\<^sub>2 :: "'bo set" and bot\<^sub>2 le\<^sub>2 ls\<^sub>2 begin sublocale bot_pair_ow .. sublocale order_pair_ow .. end locale top_ow = fixes U :: "'ao set" and top (\\\<^sub>o\<^sub>w\) assumes top_closed[simp]: "\\<^sub>o\<^sub>w \ U" begin notation top (\\\<^sub>o\<^sub>w\) end locale top_pair_ow = ord\<^sub>1: top_ow U\<^sub>1 top\<^sub>1 + ord\<^sub>2: top_ow U\<^sub>2 top\<^sub>2 for U\<^sub>1 :: "'ao set" and top\<^sub>1 and U\<^sub>2 :: "'bo set" and top\<^sub>2 begin notation top\<^sub>1 (\\\<^sub>o\<^sub>w\<^sub>.\<^sub>1\) notation top\<^sub>2 (\\\<^sub>o\<^sub>w\<^sub>.\<^sub>2\) end locale order_top_ow = order_ow U le ls + top_ow U top for U :: "'ao set" and le ls top + assumes top_greatest: "a \ U \ a \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w" begin sublocale top: ordering_top_ow U \(\\<^sub>o\<^sub>w)\ \(<\<^sub>o\<^sub>w)\ \\\<^sub>o\<^sub>w\ apply unfold_locales subgoal by simp subgoal by (simp add: top_greatest) done no_notation le (infix "\<^bold>\\<^sub>o\<^sub>w" 50) and ls (infix "\<^bold><\<^sub>o\<^sub>w" 50) and top ("\<^bold>\\<^sub>o\<^sub>w") end locale order_top_pair_ow = ord\<^sub>1: order_top_ow U\<^sub>1 le\<^sub>1 ls\<^sub>1 top\<^sub>1 + ord\<^sub>2: order_top_ow U\<^sub>2 le\<^sub>2 ls\<^sub>2 top\<^sub>2 for U\<^sub>1 :: "'ao set" and le\<^sub>1 ls\<^sub>1 top\<^sub>1 and U\<^sub>2 :: "'bo set" and le\<^sub>2 ls\<^sub>2 top\<^sub>2 begin sublocale top_pair_ow .. sublocale order_pair_ow .. end subsubsection\Transfer rules\ context includes lifting_syntax begin lemma ordering_top_transfer[transfer_rule]: assumes [transfer_rule]: "bi_unique A" "right_total A" shows "((A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> A ===> (=)) (ordering_top_ow (Collect (Domainp A))) ordering_top" proof- let ?P = "((A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> A ===> (=))" let ?ordering_top_ow = "ordering_top_ow (Collect (Domainp A))" have "?P ?ordering_top_ow (\le ls ext. ext \ UNIV \ ordering_top le ls ext)" unfolding ordering_top_ow_def ordering_top_def ordering_top_ow_axioms_def ordering_top_axioms_def apply transfer_prover_start apply transfer_step+ by blast thus ?thesis by simp qed lemma order_bot_transfer[transfer_rule]: assumes [transfer_rule]: "bi_unique A" "right_total A" shows "(A ===> (A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> (=)) (order_bot_ow (Collect (Domainp A))) class.order_bot" proof- let ?P = "(A ===> (A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> (=))" let ?order_bot_ow = "order_bot_ow (Collect (Domainp A))" have "?P ?order_bot_ow (\bot le ls. bot \ UNIV \ class.order_bot bot le ls)" unfolding class.order_bot_def order_bot_ow_def class.order_bot_axioms_def order_bot_ow_axioms_def bot_ow_def apply transfer_prover_start apply transfer_step+ by blast thus ?thesis by simp qed lemma order_top_transfer[transfer_rule]: assumes [transfer_rule]: "bi_unique A" "right_total A" shows "((A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> A ===> (=)) (order_top_ow (Collect (Domainp A))) class.order_top" proof- let ?P = "((A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> A ===> (=))" let ?order_top_ow = "order_top_ow (Collect (Domainp A))" have "?P ?order_top_ow (\le ls top. top \ UNIV \ class.order_top le ls top)" unfolding class.order_top_def order_top_ow_def class.order_top_axioms_def order_top_ow_axioms_def top_ow_def apply transfer_prover_start apply transfer_step+ by blast thus ?thesis by simp qed end subsubsection\Relativization\ context ordering_top_ow begin tts_context tts: (?'a to U) rewriting ctr_simps substituting ordering_top_ow_axioms eliminating through simp applying [OF top_closed] begin tts_lemma extremum_uniqueI: assumes "\<^bold>\\<^sub>o\<^sub>w \<^bold>\\<^sub>o\<^sub>w \<^bold>\\<^sub>o\<^sub>w" shows "\<^bold>\\<^sub>o\<^sub>w = \<^bold>\\<^sub>o\<^sub>w" is ordering_top.extremum_uniqueI. tts_lemma extremum_unique: shows "(\<^bold>\\<^sub>o\<^sub>w \<^bold>\\<^sub>o\<^sub>w \<^bold>\\<^sub>o\<^sub>w) = (\<^bold>\\<^sub>o\<^sub>w = \<^bold>\\<^sub>o\<^sub>w)" is ordering_top.extremum_unique. tts_lemma extremum_strict: shows "\ \<^bold>\\<^sub>o\<^sub>w \<^bold><\<^sub>o\<^sub>w \<^bold>\\<^sub>o\<^sub>w" is ordering_top.extremum_strict. tts_lemma not_eq_extremum: shows "(\<^bold>\\<^sub>o\<^sub>w \ \<^bold>\\<^sub>o\<^sub>w) = (\<^bold>\\<^sub>o\<^sub>w \<^bold><\<^sub>o\<^sub>w \<^bold>\\<^sub>o\<^sub>w)" is ordering_top.not_eq_extremum. end end context order_bot_ow begin tts_context tts: (?'a to U) rewriting ctr_simps substituting order_bot_ow_axioms eliminating through simp begin tts_lemma bdd_above_bot: assumes "A \ U" shows "bdd_below A" is order_bot_class.bdd_below_bot. tts_lemma not_less_bot: assumes "a \ U" shows "\ a <\<^sub>o\<^sub>w \\<^sub>o\<^sub>w" is order_bot_class.not_less_bot. tts_lemma max_bot: assumes "x \ U" shows "max \\<^sub>o\<^sub>w x = x" is order_bot_class.max_bot. tts_lemma max_bot2: assumes "x \ U" shows "max x \\<^sub>o\<^sub>w = x" is order_bot_class.max_bot2. tts_lemma min_bot: assumes "x \ U" shows "min \\<^sub>o\<^sub>w x = \\<^sub>o\<^sub>w" is order_bot_class.min_bot. tts_lemma min_bot2: assumes "x \ U" shows "min x \\<^sub>o\<^sub>w = \\<^sub>o\<^sub>w" is order_bot_class.min_bot2. tts_lemma bot_unique: assumes "a \ U" shows "(a \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w) = (a = \\<^sub>o\<^sub>w)" is order_bot_class.bot_unique. tts_lemma bot_less: assumes "a \ U" shows "(a \ \\<^sub>o\<^sub>w) = (\\<^sub>o\<^sub>w <\<^sub>o\<^sub>w a)" is order_bot_class.bot_less. tts_lemma atLeast_eq_UNIV_iff: assumes "x \ U" shows "({x..\<^sub>o\<^sub>w} = U) = (x = \\<^sub>o\<^sub>w)" is order_bot_class.atLeast_eq_UNIV_iff. tts_lemma le_bot: assumes "a \ U" and "a \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w" shows "a = \\<^sub>o\<^sub>w" is order_bot_class.le_bot. end end context order_top_ow begin tts_context tts: (?'a to U) rewriting ctr_simps substituting order_top_ow_axioms eliminating through simp begin tts_lemma bdd_above_top: assumes "A \ U" shows "bdd_above A" is order_top_class.bdd_above_top. tts_lemma not_top_less: assumes "a \ U" shows "\ \\<^sub>o\<^sub>w <\<^sub>o\<^sub>w a" is order_top_class.not_top_less. tts_lemma max_top: assumes "x \ U" shows "max \\<^sub>o\<^sub>w x = \\<^sub>o\<^sub>w" is order_top_class.max_top. tts_lemma max_top2: assumes "x \ U" shows "max x \\<^sub>o\<^sub>w = \\<^sub>o\<^sub>w" is order_top_class.max_top2. tts_lemma min_top: assumes "x \ U" shows "min \\<^sub>o\<^sub>w x = x" is order_top_class.min_top. tts_lemma min_top2: assumes "x \ U" shows "min x \\<^sub>o\<^sub>w = x" is order_top_class.min_top2. tts_lemma top_unique: assumes "a \ U" shows "(\\<^sub>o\<^sub>w \\<^sub>o\<^sub>w a) = (a = \\<^sub>o\<^sub>w)" is order_top_class.top_unique. tts_lemma less_top: assumes "a \ U" shows "(a \ \\<^sub>o\<^sub>w) = (a <\<^sub>o\<^sub>w \\<^sub>o\<^sub>w)" is order_top_class.less_top. tts_lemma atMost_eq_UNIV_iff: assumes "x \ U" shows "({..\<^sub>o\<^sub>wx} = U) = (x = \\<^sub>o\<^sub>w)" is order_top_class.atMost_eq_UNIV_iff. tts_lemma top_le: assumes "a \ U" and "\\<^sub>o\<^sub>w \\<^sub>o\<^sub>w a" shows "a = \\<^sub>o\<^sub>w" is order_top_class.top_le. end end text\\newpage\ end \ No newline at end of file diff --git a/thys/WOOT_Strong_Eventual_Consistency/Sorting.thy b/thys/WOOT_Strong_Eventual_Consistency/Sorting.thy --- a/thys/WOOT_Strong_Eventual_Consistency/Sorting.thy +++ b/thys/WOOT_Strong_Eventual_Consistency/Sorting.thy @@ -1,141 +1,141 @@ subsection \Sorting\ text \Some preliminary lemmas about sorting.\ theory Sorting imports Main "HOL.List" "HOL-Library.Sublist" begin lemma insort: assumes "Suc l < length s" assumes "s ! l < (v :: 'a :: linorder)" assumes "s ! (l+1) > v" assumes "sorted_wrt (<) s" shows "sorted_wrt (<) ((take (Suc l) s)@v#(drop (Suc l) s))" proof - have "sorted_wrt (<) (take (Suc l) s@(drop (Suc l) s))" using assms(4) by simp moreover have "\x. x \ set (take (Suc l) s) = (\i. i < (Suc l) \ i < length s \ s ! i = x)" by (metis in_set_conv_nth length_take min_less_iff_conj nth_take) hence "\x. x \ set (take (Suc l) s) \ x < v" using assms apply (simp) using less_Suc_eq sorted_wrt_nth_less by fastforce moreover have "\x. x \ set (drop (Suc l) s) = (\i. Suc l + i < length s \ s ! (Suc l + i) = x)" using assms(1) by (simp add:in_set_conv_nth add.commute less_diff_conv) hence "\x. x \ set (drop (Suc l) s) \ x > v" using assms apply (simp) by (metis add.right_neutral add_diff_cancel_left' diff_Suc_Suc diff_is_0_eq' leI le_less_trans less_imp_le sorted_wrt_iff_nth_less) ultimately show ?thesis by (simp add:sorted_wrt_append del:append_take_drop_id) qed lemma sorted_wrt_irrefl_distinct: assumes "irreflp r" shows "sorted_wrt r xs \ distinct xs" using assms by (induction xs, simp, simp, meson irreflp_def) lemma sort_set_unique_h: assumes "irreflp r \ transp r" assumes "set (x#xs) = set (y#ys)" assumes "\z \ set xs. r x z" assumes "\z \ set ys. r y z" shows "x = y \ set xs = set ys" by (metis assms insert_eq_iff irreflp_def list.set_intros(1) list.simps(15) set_ConsD transpD) lemma sort_set_unique_rel: assumes "irreflp r \ transp r" assumes "set x = set y" assumes "sorted_wrt r x" assumes "sorted_wrt r y" shows "x = y" proof - have "length x = length y" using assms by (metis sorted_wrt_irrefl_distinct distinct_card) then show ?thesis using assms apply(induct rule:list_induct2, simp, simp) by (metis assms(1) list.simps(15) sort_set_unique_h) qed lemma sort_set_unique: assumes "set x = set y" assumes "sorted_wrt (<) (map (f :: ('a \ ('b :: linorder))) x)" assumes "sorted_wrt (<) (map f y)" shows "x = y" using assms apply (simp add:sorted_wrt_map) by (metis (no_types, lifting) irreflp_def less_irrefl sort_set_unique_rel - transpD transpI transp_less) + transpD transpI transp_on_less) text \If two sequences contain the same element and strictly increasing with respect.\ lemma subseq_imp_sorted: assumes "subseq s t" assumes "sorted_wrt p t" shows "sorted_wrt p s" proof - have "sorted_wrt p s \ \ sorted_wrt p t" apply (rule list_emb.induct[where P="(=)"]) using list_emb_set assms by fastforce+ thus ?thesis using assms by blast qed text \If a sequence @{text t} is sorted with respect to a relation @{text p} then a subsequence will be as well.\ fun to_ord where "to_ord r x y = (\(r\<^sup>*\<^sup>* y x))" lemma trancl_idemp: "r\<^sup>+\<^sup>+\<^sup>+\<^sup>+ x y = r\<^sup>+\<^sup>+ x y" by (metis r_into_rtranclp reflclp_tranclp rtranclp_idemp rtranclp_reflclp rtranclp_tranclp_tranclp tranclp.cases tranclp.r_into_trancl) lemma top_sort: fixes rp assumes "acyclicP r" shows "finite s \ (\l. set l = s \ sorted_wrt (to_ord r) l \ distinct l)" proof (induction "card s" arbitrary:s) case 0 then show ?case by auto next case (Suc n) hence "s \ {}" by auto moreover have "acyclicP (r\<^sup>+\<^sup>+)" using assms by (simp add:acyclic_def trancl_def trancl_idemp) hence "acyclic ({(x,y). r\<^sup>+\<^sup>+ x y} \ s \ s)" by (meson acyclic_subset inf_le1) hence "wf ({(x,y). r\<^sup>+\<^sup>+ x y} \ s \ s)" using Suc by (metis card.infinite finite_Int finite_SigmaI nat.distinct(1) wf_iff_acyclic_if_finite) ultimately obtain z where "z \ s \ (\y. (y, z) \ ({(x,y). r\<^sup>+\<^sup>+ x y} \ s \ s) \ y \ s)" by (metis ex_in_conv wf_eq_minimal) hence z_def: "z \ s \ (\y. r\<^sup>+\<^sup>+ y z \ y \ s)" by blast hence "card (s - {z}) = n" by (metis One_nat_def Suc.hyps(2) card_Diff_singleton_if card.infinite diff_Suc_Suc diff_zero nat.simps(3)) then obtain l where l_def: "set l = s - {z} \ sorted_wrt (to_ord r) l \ distinct l" by (metis Zero_not_Suc card.infinite finite_Diff Suc) hence "set (z#l) = s" using z_def by auto moreover have "\y \ set l. \(r\<^sup>*\<^sup>* y z)" using z_def l_def rtranclpD by force ultimately show ?case by (metis distinct.simps(2) insert_absorb l_def list.simps(15) sorted_wrt.simps(2) to_ord.elims(3)) qed lemma top_sort_eff: assumes "irreflp p\<^sup>+\<^sup>+" assumes "sorted_wrt (to_ord p) x" assumes "i < length x" assumes "j < length x" assumes "(p\<^sup>+\<^sup>+ (x ! i) (x ! j))" shows "i < j" using assms apply (cases "i > j") apply (metis sorted_wrt_nth_less r_into_rtranclp reflclp_tranclp rtranclp_idemp rtranclp_reflclp to_ord.simps) by (metis irreflp_def nat_neq_iff) end \ No newline at end of file