src/HOL/IMP/Abs_Int1_parity.thy
author wenzelm
Thu, 08 Jun 2017 12:54:55 +0200
changeset 66037 58d2e41afbfe
parent 64593 50c715579715
child 67406 23307fd33906
permissions -rw-r--r--
more official session qualifier;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
     1
(* Author: Tobias Nipkow *)
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
     2
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     3
theory Abs_Int1_parity
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     4
imports Abs_Int1
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     5
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     6
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     7
subsection "Parity Analysis"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     8
58310
91ea607a34d8 updated news
blanchet
parents: 58249
diff changeset
     9
datatype parity = Even | Odd | Either
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    10
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    11
text{* Instantiation of class @{class order} with type @{typ parity}: *}
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    12
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    13
instantiation parity :: order
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    14
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    15
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    16
text{* First the definition of the interface function @{text"\<le>"}. Note that
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    17
the header of the definition must refer to the ascii name @{const less_eq} of the
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    18
constants as @{text less_eq_parity} and the definition is named @{text
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    19
less_eq_parity_def}.  Inside the definition the symbolic names can be used. *}
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    20
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    21
definition less_eq_parity where
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    22
"x \<le> y = (y = Either \<or> x=y)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    23
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    24
text{* We also need @{text"<"}, which is defined canonically: *}
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    25
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    26
definition less_parity where
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    27
"x < y = (x \<le> y \<and> \<not> y \<le> (x::parity))"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    28
51625
bd3358aac5d2 tuned document
nipkow
parents: 51624
diff changeset
    29
text{*\noindent(The type annotation is necessary to fix the type of the polymorphic predicates.)
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    30
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    31
Now the instance proof, i.e.\ the proof that the definition fulfills
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    32
the axioms (assumptions) of the class. The initial proof-step generates the
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    33
necessary proof obligations. *}
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    34
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    35
instance
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    36
proof
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    37
  fix x::parity show "x \<le> x" by(auto simp: less_eq_parity_def)
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    38
next
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    39
  fix x y z :: parity assume "x \<le> y" "y \<le> z" thus "x \<le> z"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    40
    by(auto simp: less_eq_parity_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    41
next
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    42
  fix x y :: parity assume "x \<le> y" "y \<le> x" thus "x = y"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    43
    by(auto simp: less_eq_parity_def)
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    44
next
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    45
  fix x y :: parity show "(x < y) = (x \<le> y \<and> \<not> y \<le> x)" by(rule less_parity_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    46
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    47
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    48
end
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    49
51826
054a40461449 canonical names of classes
nipkow
parents: 51711
diff changeset
    50
text{* Instantiation of class @{class semilattice_sup_top} with type @{typ parity}: *}
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    51
51826
054a40461449 canonical names of classes
nipkow
parents: 51711
diff changeset
    52
instantiation parity :: semilattice_sup_top
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    53
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    54
51389
8a9f0503b1c0 factored out Step
nipkow
parents: 51372
diff changeset
    55
definition sup_parity where
51624
nipkow
parents: 51389
diff changeset
    56
"x \<squnion> y = (if x = y then x else Either)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    57
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    58
definition top_parity where
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    59
"\<top> = Either"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    60
61179
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
    61
text{* Now the instance proof. This time we take a shortcut with the help of
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
    62
proof method @{text goal_cases}: it creates cases 1 ... n for the subgoals
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
    63
1 ... n; in case i, i is also the name of the assumptions of subgoal i and
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
    64
@{text "case?"} refers to the conclusion of subgoal i.
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
    65
The class axioms are presented in the same order as in the class definition. *}
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    66
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    67
instance
61179
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
    68
proof (standard, goal_cases)
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
    69
  case 1 (*sup1*) show ?case by(auto simp: less_eq_parity_def sup_parity_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    70
next
61179
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
    71
  case 2 (*sup2*) show ?case by(auto simp: less_eq_parity_def sup_parity_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    72
next
61179
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
    73
  case 3 (*sup least*) thus ?case by(auto simp: less_eq_parity_def sup_parity_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    74
next
61179
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
    75
  case 4 (*top*) show ?case by(auto simp: less_eq_parity_def top_parity_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    76
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    77
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    78
end
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    79
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    80
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    81
text{* Now we define the functions used for instantiating the abstract
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    82
interpretation locales. Note that the Isabelle terminology is
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    83
\emph{interpretation}, not \emph{instantiation} of locales, but we use
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    84
instantiation to avoid confusion with abstract interpretation.  *}
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    85
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    86
fun \<gamma>_parity :: "parity \<Rightarrow> val set" where
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    87
"\<gamma>_parity Even = {i. i mod 2 = 0}" |
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    88
"\<gamma>_parity Odd  = {i. i mod 2 = 1}" |
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    89
"\<gamma>_parity Either = UNIV"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    90
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    91
fun num_parity :: "val \<Rightarrow> parity" where
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    92
"num_parity i = (if i mod 2 = 0 then Even else Odd)"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    93
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    94
fun plus_parity :: "parity \<Rightarrow> parity \<Rightarrow> parity" where
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    95
"plus_parity Even Even = Even" |
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    96
"plus_parity Odd  Odd  = Even" |
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    97
"plus_parity Even Odd  = Odd" |
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    98
"plus_parity Odd  Even = Odd" |
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    99
"plus_parity Either y  = Either" |
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   100
"plus_parity x Either  = Either"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   101
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   102
text{* First we instantiate the abstract value interface and prove that the
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   103
functions on type @{typ parity} have all the necessary properties: *}
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   104
61890
f6ded81f5690 abandoned attempt to unify sublocale and interpretation into global theories
haftmann
parents: 61671
diff changeset
   105
global_interpretation Val_semilattice
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   106
where \<gamma> = \<gamma>_parity and num' = num_parity and plus' = plus_parity
61179
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
   107
proof (standard, goal_cases) txt{* subgoals are the locale axioms *}
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
   108
  case 1 thus ?case by(auto simp: less_eq_parity_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   109
next
61179
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
   110
  case 2 show ?case by(auto simp: top_parity_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   111
next
61179
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
   112
  case 3 show ?case by auto
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
   113
next
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
   114
  case (4 _ a1 _ a2) thus ?case
64593
50c715579715 reoriented congruence rules in non-explosive direction
haftmann
parents: 61890
diff changeset
   115
    by (induction a1 a2 rule: plus_parity.induct)
50c715579715 reoriented congruence rules in non-explosive direction
haftmann
parents: 61890
diff changeset
   116
      (auto simp add: mod_add_eq [symmetric])
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   117
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   118
61179
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
   119
text{* In case 4 we needed to refer to particular variables.
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
   120
Writing (i x y z) fixes the names of the variables in case i to be x, y and z
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
   121
in the left-to-right order in which the variables occur in the subgoal.
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
   122
Underscores are anonymous placeholders for variable names we don't care to fix. *}
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
   123
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   124
text{* Instantiating the abstract interpretation locale requires no more
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   125
proofs (they happened in the instatiation above) but delivers the
49344
nipkow
parents: 49188
diff changeset
   126
instantiated abstract interpreter which we call @{text AI_parity}: *}
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   127
61890
f6ded81f5690 abandoned attempt to unify sublocale and interpretation into global theories
haftmann
parents: 61671
diff changeset
   128
global_interpretation Abs_Int
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   129
where \<gamma> = \<gamma>_parity and num' = num_parity and plus' = plus_parity
61671
20d4cd2ceab2 prefer "rewrites" and "defines" to note rewrite morphisms
haftmann
parents: 61179
diff changeset
   130
defines aval_parity = aval' and step_parity = step' and AI_parity = AI
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   131
..
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   132
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   133
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   134
subsubsection "Tests"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   135
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   136
definition "test1_parity =
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 51826
diff changeset
   137
  ''x'' ::= N 1;;
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   138
  WHILE Less (V ''x'') (N 100) DO ''x'' ::= Plus (V ''x'') (N 2)"
56927
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55600
diff changeset
   139
value "show_acom (the(AI_parity test1_parity))"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   140
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   141
definition "test2_parity =
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 51826
diff changeset
   142
  ''x'' ::= N 1;;
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   143
  WHILE Less (V ''x'') (N 100) DO ''x'' ::= Plus (V ''x'') (N 3)"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   144
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51625
diff changeset
   145
definition "steps c i = ((step_parity \<top>) ^^ i) (bot c)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   146
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   147
value "show_acom (steps test2_parity 0)"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   148
value "show_acom (steps test2_parity 1)"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   149
value "show_acom (steps test2_parity 2)"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   150
value "show_acom (steps test2_parity 3)"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   151
value "show_acom (steps test2_parity 4)"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   152
value "show_acom (steps test2_parity 5)"
49188
22f7e7b68f50 adjusted examples
nipkow
parents: 47613
diff changeset
   153
value "show_acom (steps test2_parity 6)"
50995
nipkow
parents: 49433
diff changeset
   154
value "show_acom (the(AI_parity test2_parity))"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   155
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   156
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   157
subsubsection "Termination"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   158
61890
f6ded81f5690 abandoned attempt to unify sublocale and interpretation into global theories
haftmann
parents: 61671
diff changeset
   159
global_interpretation Abs_Int_mono
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   160
where \<gamma> = \<gamma>_parity and num' = num_parity and plus' = plus_parity
61179
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
   161
proof (standard, goal_cases)
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
   162
  case (1 _ a1 _ a2) thus ?case
52525
1e09c034d6c3 tund proofs
nipkow
parents: 52504
diff changeset
   163
    by(induction a1 a2 rule: plus_parity.induct)
1e09c034d6c3 tund proofs
nipkow
parents: 52504
diff changeset
   164
      (auto simp add:less_eq_parity_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   165
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   166
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   167
definition m_parity :: "parity \<Rightarrow> nat" where
51389
8a9f0503b1c0 factored out Step
nipkow
parents: 51372
diff changeset
   168
"m_parity x = (if x = Either then 0 else 1)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   169
61890
f6ded81f5690 abandoned attempt to unify sublocale and interpretation into global theories
haftmann
parents: 61671
diff changeset
   170
global_interpretation Abs_Int_measure
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   171
where \<gamma> = \<gamma>_parity and num' = num_parity and plus' = plus_parity
49433
nipkow
parents: 49399
diff changeset
   172
and m = m_parity and h = "1"
61179
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
   173
proof (standard, goal_cases)
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
   174
  case 1 thus ?case by(auto simp add: m_parity_def less_eq_parity_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   175
next
61179
16775cad1a5c goali -> i
nipkow
parents: 58310
diff changeset
   176
  case 2 thus ?case by(auto simp add: m_parity_def less_eq_parity_def less_parity_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   177
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   178
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   179
thm AI_Some_measure
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   180
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   181
end