| 53430 |      1 | theory Simps_Case_Conv_Examples imports
 | 
|  |      2 |   "~~/src/HOL/Library/Simps_Case_Conv"
 | 
|  |      3 | begin
 | 
|  |      4 | 
 | 
|  |      5 | section {* Tests for the Simps<->Case conversion tools *}
 | 
|  |      6 | 
 | 
|  |      7 | fun foo where
 | 
|  |      8 |   "foo (x # xs) Nil = 0" |
 | 
|  |      9 |   "foo (x # xs) (y # ys) = foo [] []" |
 | 
|  |     10 |   "foo Nil (y # ys) = 1" |
 | 
|  |     11 |   "foo Nil Nil = 3"
 | 
|  |     12 | 
 | 
|  |     13 | fun bar where
 | 
|  |     14 |   "bar x 0 y = 0 + x" |
 | 
|  |     15 |   "bar x (Suc n) y = n + x"
 | 
|  |     16 | 
 | 
|  |     17 | definition
 | 
|  |     18 |   split_rule_test :: "((nat => 'a) + ('b * (('b => 'a) option))) => ('a => nat) => nat"
 | 
|  |     19 | where
 | 
|  |     20 |  "split_rule_test x f = f (case x of Inl af \<Rightarrow> af 1
 | 
|  |     21 |     | Inr (b, None) => inv f 0
 | 
|  |     22 |     | Inr (b, Some g) => g b)"
 | 
|  |     23 | 
 | 
|  |     24 | definition test where "test x y = (case x of None => (case y of [] => 1 | _ # _ => 2) | Some x => x)"
 | 
|  |     25 | 
 | 
|  |     26 | definition nosplit where "nosplit x = x @ (case x of [] \<Rightarrow> [1] | xs \<Rightarrow> xs)"
 | 
|  |     27 | 
 | 
|  |     28 | 
 | 
|  |     29 | text {* Function with complete, non-overlapping patterns *}
 | 
|  |     30 | case_of_simps foo_cases1: foo.simps
 | 
|  |     31 | 
 | 
|  |     32 | text {* Redundant equations are ignored *}
 | 
|  |     33 | case_of_simps foo_cases2: foo.simps foo.simps
 | 
|  |     34 | print_theorems
 | 
|  |     35 | 
 | 
|  |     36 | text {* Variable patterns *}
 | 
|  |     37 | case_of_simps bar_cases: bar.simps
 | 
|  |     38 | print_theorems
 | 
|  |     39 | 
 | 
|  |     40 | text {* Case expression not at top level *}
 | 
|  |     41 | simps_of_case split_rule_test_simps: split_rule_test_def
 | 
|  |     42 | print_theorems
 | 
|  |     43 | 
 | 
|  |     44 | text {* Argument occurs both as case parameter and seperately*}
 | 
|  |     45 | simps_of_case nosplit_simps1: nosplit_def
 | 
|  |     46 | print_theorems
 | 
|  |     47 | 
 | 
|  |     48 | text {* Nested case expressions *}
 | 
|  |     49 | simps_of_case test_simps1: test_def
 | 
|  |     50 | print_theorems
 | 
|  |     51 | 
 | 
|  |     52 | text {* Partial split of case *}
 | 
|  |     53 | simps_of_case nosplit_simps2: nosplit_def (splits: list.split)
 | 
|  |     54 | print_theorems
 | 
|  |     55 | simps_of_case test_simps2: test_def (splits: option.split)
 | 
|  |     56 | print_theorems
 | 
|  |     57 | 
 | 
|  |     58 | text {* Reversal *}
 | 
|  |     59 | case_of_simps test_def1: test_simps1
 | 
|  |     60 | print_theorems
 | 
|  |     61 | 
 | 
|  |     62 | text {* Case expressions on RHS *}
 | 
|  |     63 | case_of_simps test_def2: test_simps2
 | 
|  |     64 | print_theorems
 | 
|  |     65 | 
 | 
|  |     66 | text {* Partial split of simps *}
 | 
|  |     67 | case_of_simps foo_cons_def: foo.simps(1,2)
 | 
|  |     68 | print_theorems
 | 
|  |     69 | 
 | 
|  |     70 | 
 | 
|  |     71 | end
 |