| author | nipkow | 
| Thu, 01 Apr 2010 09:31:58 +0200 | |
| changeset 36070 | d80e5d3c8fe1 | 
| parent 36055 | 537876d0fa62 | 
| child 36176 | 3fe7e97ccca8 | 
| permissions | -rw-r--r-- | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1 | theory Predicate_Compile_Examples | 
| 35954 | 2 | imports Predicate_Compile_Alternative_Defs | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 3 | begin | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 4 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 5 | subsection {* Basic predicates *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 6 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 7 | inductive False' :: "bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 8 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 9 | code_pred (expected_modes: bool) False' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 10 | code_pred [dseq] False' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 11 | code_pred [random_dseq] False' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 12 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 13 | values [expected "{}" pred] "{x. False'}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 14 | values [expected "{}" dseq 1] "{x. False'}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 15 | values [expected "{}" random_dseq 1, 1, 1] "{x. False'}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 16 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 17 | value "False'" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 18 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 19 | inductive True' :: "bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 20 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 21 | "True ==> True'" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 22 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 23 | code_pred True' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 24 | code_pred [dseq] True' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 25 | code_pred [random_dseq] True' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 26 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 27 | thm True'.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 28 | thm True'.dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 29 | thm True'.random_dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 30 | values [expected "{()}" ]"{x. True'}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 31 | values [expected "{}" dseq 0] "{x. True'}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 32 | values [expected "{()}" dseq 1] "{x. True'}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 33 | values [expected "{()}" dseq 2] "{x. True'}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 34 | values [expected "{}" random_dseq 1, 1, 0] "{x. True'}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 35 | values [expected "{}" random_dseq 1, 1, 1] "{x. True'}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 36 | values [expected "{()}" random_dseq 1, 1, 2] "{x. True'}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 37 | values [expected "{()}" random_dseq 1, 1, 3] "{x. True'}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 38 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 39 | inductive EmptySet :: "'a \<Rightarrow> bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 40 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 41 | code_pred (expected_modes: o => bool, i => bool) EmptySet . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 42 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 43 | definition EmptySet' :: "'a \<Rightarrow> bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 44 | where "EmptySet' = {}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 45 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 46 | code_pred (expected_modes: o => bool, i => bool) [inductify] EmptySet' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 47 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 48 | inductive EmptyRel :: "'a \<Rightarrow> 'b \<Rightarrow> bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 49 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 50 | code_pred (expected_modes: o => o => bool, i => o => bool, o => i => bool, i => i => bool) EmptyRel . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 51 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 52 | inductive EmptyClosure :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 53 | for r :: "'a \<Rightarrow> 'a \<Rightarrow> bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 54 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 55 | code_pred | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 56 | (expected_modes: (o => o => bool) => o => o => bool, (o => o => bool) => i => o => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 57 | (o => o => bool) => o => i => bool, (o => o => bool) => i => i => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 58 | (i => o => bool) => o => o => bool, (i => o => bool) => i => o => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 59 | (i => o => bool) => o => i => bool, (i => o => bool) => i => i => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 60 | (o => i => bool) => o => o => bool, (o => i => bool) => i => o => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 61 | (o => i => bool) => o => i => bool, (o => i => bool) => i => i => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 62 | (i => i => bool) => o => o => bool, (i => i => bool) => i => o => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 63 | (i => i => bool) => o => i => bool, (i => i => bool) => i => i => bool) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 64 | EmptyClosure . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 65 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 66 | thm EmptyClosure.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 67 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 68 | (* TODO: inductive package is broken! | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 69 | inductive False'' :: "bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 70 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 71 | "False \<Longrightarrow> False''" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 72 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 73 | code_pred (expected_modes: []) False'' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 74 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 75 | inductive EmptySet'' :: "'a \<Rightarrow> bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 76 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 77 | "False \<Longrightarrow> EmptySet'' x" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 78 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 79 | code_pred (expected_modes: [1]) EmptySet'' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 80 | code_pred (expected_modes: [], [1]) [inductify] EmptySet'' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 81 | *) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 82 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 83 | consts a' :: 'a | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 84 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 85 | inductive Fact :: "'a \<Rightarrow> 'a \<Rightarrow> bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 86 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 87 | "Fact a' a'" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 88 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 89 | code_pred (expected_modes: o => o => bool, i => o => bool, o => i => bool, i => i => bool) Fact . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 90 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 91 | inductive zerozero :: "nat * nat => bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 92 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 93 | "zerozero (0, 0)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 94 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 95 | code_pred (expected_modes: i => bool, i * o => bool, o * i => bool, o => bool) zerozero . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 96 | code_pred [dseq] zerozero . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 97 | code_pred [random_dseq] zerozero . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 98 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 99 | thm zerozero.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 100 | thm zerozero.dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 101 | thm zerozero.random_dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 102 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 103 | text {* We expect the user to expand the tuples in the values command.
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 104 | The following values command is not supported. *} | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 105 | (*values "{x. zerozero x}" *)
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 106 | text {* Instead, the user must type *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 107 | values "{(x, y). zerozero (x, y)}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 108 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 109 | values [expected "{}" dseq 0] "{(x, y). zerozero (x, y)}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 110 | values [expected "{(0::nat, 0::nat)}" dseq 1] "{(x, y). zerozero (x, y)}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 111 | values [expected "{(0::nat, 0::nat)}" dseq 2] "{(x, y). zerozero (x, y)}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 112 | values [expected "{}" random_dseq 1, 1, 2] "{(x, y). zerozero (x, y)}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 113 | values [expected "{(0::nat, 0:: nat)}" random_dseq 1, 1, 3] "{(x, y). zerozero (x, y)}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 114 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 115 | inductive nested_tuples :: "((int * int) * int * int) => bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 116 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 117 | "nested_tuples ((0, 1), 2, 3)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 118 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 119 | code_pred nested_tuples . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 120 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 121 | inductive JamesBond :: "nat => int => code_numeral => bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 122 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 123 | "JamesBond 0 0 7" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 124 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 125 | code_pred JamesBond . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 126 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 127 | values [expected "{(0::nat, 0::int , 7::code_numeral)}"] "{(a, b, c). JamesBond a b c}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 128 | values [expected "{(0::nat, 7::code_numeral, 0:: int)}"] "{(a, c, b). JamesBond a b c}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 129 | values [expected "{(0::int, 0::nat, 7::code_numeral)}"] "{(b, a, c). JamesBond a b c}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 130 | values [expected "{(0::int, 7::code_numeral, 0::nat)}"] "{(b, c, a). JamesBond a b c}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 131 | values [expected "{(7::code_numeral, 0::nat, 0::int)}"] "{(c, a, b). JamesBond a b c}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 132 | values [expected "{(7::code_numeral, 0::int, 0::nat)}"] "{(c, b, a). JamesBond a b c}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 133 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 134 | values [expected "{(7::code_numeral, 0::int)}"] "{(a, b). JamesBond 0 b a}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 135 | values [expected "{(7::code_numeral, 0::nat)}"] "{(c, a). JamesBond a 0 c}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 136 | values [expected "{(0::nat, 7::code_numeral)}"] "{(a, c). JamesBond a 0 c}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 137 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 138 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 139 | subsection {* Alternative Rules *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 140 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 141 | datatype char = C | D | E | F | G | H | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 142 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 143 | inductive is_C_or_D | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 144 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 145 | "(x = C) \<or> (x = D) ==> is_C_or_D x" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 146 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 147 | code_pred (expected_modes: i => bool) is_C_or_D . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 148 | thm is_C_or_D.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 149 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 150 | inductive is_D_or_E | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 151 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 152 | "(x = D) \<or> (x = E) ==> is_D_or_E x" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 153 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 154 | lemma [code_pred_intro]: | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 155 | "is_D_or_E D" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 156 | by (auto intro: is_D_or_E.intros) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 157 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 158 | lemma [code_pred_intro]: | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 159 | "is_D_or_E E" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 160 | by (auto intro: is_D_or_E.intros) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 161 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 162 | code_pred (expected_modes: o => bool, i => bool) is_D_or_E | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 163 | proof - | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 164 | case is_D_or_E | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 165 | from this(1) show thesis | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 166 | proof | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 167 | fix xa | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 168 | assume x: "x = xa" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 169 | assume "xa = D \<or> xa = E" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 170 | from this show thesis | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 171 | proof | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 172 | assume "xa = D" from this x is_D_or_E(2) show thesis by simp | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 173 | next | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 174 | assume "xa = E" from this x is_D_or_E(3) show thesis by simp | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 175 | qed | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 176 | qed | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 177 | qed | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 178 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 179 | thm is_D_or_E.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 180 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 181 | inductive is_F_or_G | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 182 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 183 | "x = F \<or> x = G ==> is_F_or_G x" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 184 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 185 | lemma [code_pred_intro]: | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 186 | "is_F_or_G F" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 187 | by (auto intro: is_F_or_G.intros) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 188 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 189 | lemma [code_pred_intro]: | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 190 | "is_F_or_G G" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 191 | by (auto intro: is_F_or_G.intros) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 192 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 193 | inductive is_FGH | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 194 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 195 | "is_F_or_G x ==> is_FGH x" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 196 | | "is_FGH H" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 197 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 198 | text {* Compilation of is_FGH requires elimination rule for is_F_or_G *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 199 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 200 | code_pred (expected_modes: o => bool, i => bool) is_FGH | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 201 | proof - | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 202 | case is_F_or_G | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 203 | from this(1) show thesis | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 204 | proof | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 205 | fix xa | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 206 | assume x: "x = xa" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 207 | assume "xa = F \<or> xa = G" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 208 | from this show thesis | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 209 | proof | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 210 | assume "xa = F" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 211 | from this x is_F_or_G(2) show thesis by simp | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 212 | next | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 213 | assume "xa = G" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 214 | from this x is_F_or_G(3) show thesis by simp | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 215 | qed | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 216 | qed | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 217 | qed | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 218 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 219 | subsection {* Preprocessor Inlining  *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 220 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 221 | definition "equals == (op =)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 222 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 223 | inductive zerozero' :: "nat * nat => bool" where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 224 | "equals (x, y) (0, 0) ==> zerozero' (x, y)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 225 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 226 | code_pred (expected_modes: i => bool) zerozero' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 227 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 228 | lemma zerozero'_eq: "zerozero' x == zerozero x" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 229 | proof - | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 230 | have "zerozero' = zerozero" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 231 | apply (auto simp add: mem_def) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 232 | apply (cases rule: zerozero'.cases) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 233 | apply (auto simp add: equals_def intro: zerozero.intros) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 234 | apply (cases rule: zerozero.cases) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 235 | apply (auto simp add: equals_def intro: zerozero'.intros) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 236 | done | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 237 | from this show "zerozero' x == zerozero x" by auto | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 238 | qed | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 239 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 240 | declare zerozero'_eq [code_pred_inline] | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 241 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 242 | definition "zerozero'' x == zerozero' x" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 243 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 244 | text {* if preprocessing fails, zerozero'' will not have all modes. *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 245 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 246 | code_pred (expected_modes: i * i => bool, i * o => bool, o * i => bool, o => bool) [inductify] zerozero'' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 247 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 248 | subsection {* Sets and Numerals *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 249 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 250 | definition | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 251 |   "one_or_two = {Suc 0, (Suc (Suc 0))}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 252 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 253 | code_pred [inductify] one_or_two . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 254 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 255 | code_pred [dseq] one_or_two . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 256 | code_pred [random_dseq] one_or_two . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 257 | thm one_or_two.dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 258 | values [expected "{Suc 0::nat, 2::nat}"] "{x. one_or_two x}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 259 | values [random_dseq 0,0,10] 3 "{x. one_or_two x}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 260 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 261 | inductive one_or_two' :: "nat => bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 262 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 263 | "one_or_two' 1" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 264 | | "one_or_two' 2" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 265 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 266 | code_pred one_or_two' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 267 | thm one_or_two'.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 268 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 269 | values "{x. one_or_two' x}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 270 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 271 | definition one_or_two'': | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 272 |   "one_or_two'' == {1, (2::nat)}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 273 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 274 | code_pred [inductify] one_or_two'' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 275 | thm one_or_two''.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 276 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 277 | values "{x. one_or_two'' x}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 278 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 279 | subsection {* even predicate *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 280 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 281 | inductive even :: "nat \<Rightarrow> bool" and odd :: "nat \<Rightarrow> bool" where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 282 | "even 0" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 283 | | "even n \<Longrightarrow> odd (Suc n)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 284 | | "odd n \<Longrightarrow> even (Suc n)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 285 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 286 | code_pred (expected_modes: i => bool, o => bool) even . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 287 | code_pred [dseq] even . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 288 | code_pred [random_dseq] even . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 289 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 290 | thm odd.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 291 | thm even.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 292 | thm odd.dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 293 | thm even.dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 294 | thm odd.random_dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 295 | thm even.random_dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 296 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 297 | values "{x. even 2}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 298 | values "{x. odd 2}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 299 | values 10 "{n. even n}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 300 | values 10 "{n. odd n}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 301 | values [expected "{}" dseq 2] "{x. even 6}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 302 | values [expected "{}" dseq 6] "{x. even 6}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 303 | values [expected "{()}" dseq 7] "{x. even 6}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 304 | values [dseq 2] "{x. odd 7}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 305 | values [dseq 6] "{x. odd 7}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 306 | values [dseq 7] "{x. odd 7}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 307 | values [expected "{()}" dseq 8] "{x. odd 7}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 308 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 309 | values [expected "{}" dseq 0] 8 "{x. even x}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 310 | values [expected "{0::nat}" dseq 1] 8 "{x. even x}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 311 | values [expected "{0::nat, 2}" dseq 3] 8 "{x. even x}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 312 | values [expected "{0::nat, 2}" dseq 4] 8 "{x. even x}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 313 | values [expected "{0::nat, 2, 4}" dseq 6] 8 "{x. even x}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 314 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 315 | values [random_dseq 1, 1, 0] 8 "{x. even x}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 316 | values [random_dseq 1, 1, 1] 8 "{x. even x}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 317 | values [random_dseq 1, 1, 2] 8 "{x. even x}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 318 | values [random_dseq 1, 1, 3] 8 "{x. even x}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 319 | values [random_dseq 1, 1, 6] 8 "{x. even x}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 320 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 321 | values [expected "{}" random_dseq 1, 1, 7] "{x. odd 7}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 322 | values [random_dseq 1, 1, 8] "{x. odd 7}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 323 | values [random_dseq 1, 1, 9] "{x. odd 7}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 324 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 325 | definition odd' where "odd' x == \<not> even x" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 326 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 327 | code_pred (expected_modes: i => bool) [inductify] odd' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 328 | code_pred [dseq inductify] odd' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 329 | code_pred [random_dseq inductify] odd' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 330 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 331 | values [expected "{}" dseq 2] "{x. odd' 7}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 332 | values [expected "{()}" dseq 9] "{x. odd' 7}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 333 | values [expected "{}" dseq 2] "{x. odd' 8}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 334 | values [expected "{}" dseq 10] "{x. odd' 8}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 335 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 336 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 337 | inductive is_even :: "nat \<Rightarrow> bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 338 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 339 | "n mod 2 = 0 \<Longrightarrow> is_even n" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 340 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 341 | code_pred (expected_modes: i => bool) is_even . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 342 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 343 | subsection {* append predicate *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 344 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 345 | inductive append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 346 | "append [] xs xs" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 347 | | "append xs ys zs \<Longrightarrow> append (x # xs) ys (x # zs)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 348 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 349 | code_pred (modes: i => i => o => bool as "concat", o => o => i => bool as "slice", o => i => i => bool as prefix, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 350 | i => o => i => bool as suffix, i => i => i => bool) append . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 351 | code_pred [dseq] append . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 352 | code_pred [random_dseq] append . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 353 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 354 | thm append.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 355 | thm append.dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 356 | thm append.random_dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 357 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 358 | values "{(ys, xs). append xs ys [0, Suc 0, 2]}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 359 | values "{zs. append [0, Suc 0, 2] [17, 8] zs}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 360 | values "{ys. append [0, Suc 0, 2] ys [0, Suc 0, 2, 17, 0, 5]}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 361 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 362 | values [expected "{}" dseq 0] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 363 | values [expected "{(([]::nat list), [Suc 0, 2, 3, 4, (5::nat)])}" dseq 1] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 364 | values [dseq 4] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 365 | values [dseq 6] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 366 | values [random_dseq 1, 1, 4] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 367 | values [random_dseq 1, 1, 1] 10 "{(xs, ys, zs::int list). append xs ys zs}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 368 | values [random_dseq 1, 1, 3] 10 "{(xs, ys, zs::int list). append xs ys zs}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 369 | values [random_dseq 3, 1, 3] 10 "{(xs, ys, zs::int list). append xs ys zs}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 370 | values [random_dseq 1, 3, 3] 10 "{(xs, ys, zs::int list). append xs ys zs}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 371 | values [random_dseq 1, 1, 4] 10 "{(xs, ys, zs::int list). append xs ys zs}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 372 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 373 | value [code] "Predicate.the (concat [0::int, 1, 2] [3, 4, 5])" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 374 | value [code] "Predicate.the (slice ([]::int list))" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 375 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 376 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 377 | text {* tricky case with alternative rules *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 378 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 379 | inductive append2 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 380 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 381 | "append2 [] xs xs" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 382 | | "append2 xs ys zs \<Longrightarrow> append2 (x # xs) ys (x # zs)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 383 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 384 | lemma append2_Nil: "append2 [] (xs::'b list) xs" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 385 | by (simp add: append2.intros(1)) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 386 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 387 | lemmas [code_pred_intro] = append2_Nil append2.intros(2) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 388 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 389 | code_pred (expected_modes: i => i => o => bool, o => o => i => bool, o => i => i => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 390 | i => o => i => bool, i => i => i => bool) append2 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 391 | proof - | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 392 | case append2 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 393 | from append2(1) show thesis | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 394 | proof | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 395 | fix xs | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 396 | assume "xa = []" "xb = xs" "xc = xs" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 397 | from this append2(2) show thesis by simp | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 398 | next | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 399 | fix xs ys zs x | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 400 | assume "xa = x # xs" "xb = ys" "xc = x # zs" "append2 xs ys zs" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 401 | from this append2(3) show thesis by fastsimp | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 402 | qed | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 403 | qed | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 404 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 405 | inductive tupled_append :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 406 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 407 | "tupled_append ([], xs, xs)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 408 | | "tupled_append (xs, ys, zs) \<Longrightarrow> tupled_append (x # xs, ys, x # zs)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 409 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 410 | code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 411 | i * o * i => bool, i * i * i => bool) tupled_append . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 412 | code_pred [random_dseq] tupled_append . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 413 | thm tupled_append.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 414 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 415 | values "{xs. tupled_append ([(1::nat), 2, 3], [4, 5], xs)}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 416 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 417 | inductive tupled_append' | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 418 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 419 | "tupled_append' ([], xs, xs)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 420 | | "[| ys = fst (xa, y); x # zs = snd (xa, y); | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 421 | tupled_append' (xs, ys, zs) |] ==> tupled_append' (x # xs, xa, y)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 422 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 423 | code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 424 | i * o * i => bool, i * i * i => bool) tupled_append' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 425 | thm tupled_append'.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 426 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 427 | inductive tupled_append'' :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 428 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 429 | "tupled_append'' ([], xs, xs)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 430 | | "ys = fst yszs ==> x # zs = snd yszs ==> tupled_append'' (xs, ys, zs) \<Longrightarrow> tupled_append'' (x # xs, yszs)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 431 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 432 | code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 433 | i * o * i => bool, i * i * i => bool) tupled_append'' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 434 | thm tupled_append''.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 435 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 436 | inductive tupled_append''' :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 437 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 438 | "tupled_append''' ([], xs, xs)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 439 | | "yszs = (ys, zs) ==> tupled_append''' (xs, yszs) \<Longrightarrow> tupled_append''' (x # xs, ys, x # zs)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 440 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 441 | code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 442 | i * o * i => bool, i * i * i => bool) tupled_append''' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 443 | thm tupled_append'''.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 444 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 445 | subsection {* map_ofP predicate *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 446 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 447 | inductive map_ofP :: "('a \<times> 'b) list \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 448 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 449 | "map_ofP ((a, b)#xs) a b" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 450 | | "map_ofP xs a b \<Longrightarrow> map_ofP (x#xs) a b" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 451 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 452 | code_pred (expected_modes: i => o => o => bool, i => i => o => bool, i => o => i => bool, i => i => i => bool) map_ofP . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 453 | thm map_ofP.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 454 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 455 | subsection {* filter predicate *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 456 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 457 | inductive filter1 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 458 | for P | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 459 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 460 | "filter1 P [] []" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 461 | | "P x ==> filter1 P xs ys ==> filter1 P (x#xs) (x#ys)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 462 | | "\<not> P x ==> filter1 P xs ys ==> filter1 P (x#xs) ys" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 463 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 464 | code_pred (expected_modes: (i => bool) => i => o => bool, (i => bool) => i => i => bool) filter1 . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 465 | code_pred [dseq] filter1 . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 466 | code_pred [random_dseq] filter1 . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 467 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 468 | thm filter1.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 469 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 470 | values [expected "{[0::nat, 2, 4]}"] "{xs. filter1 even [0, 1, 2, 3, 4] xs}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 471 | values [expected "{}" dseq 9] "{xs. filter1 even [0, 1, 2, 3, 4] xs}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 472 | values [expected "{[0::nat, 2, 4]}" dseq 10] "{xs. filter1 even [0, 1, 2, 3, 4] xs}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 473 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 474 | inductive filter2 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 475 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 476 | "filter2 P [] []" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 477 | | "P x ==> filter2 P xs ys ==> filter2 P (x#xs) (x#ys)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 478 | | "\<not> P x ==> filter2 P xs ys ==> filter2 P (x#xs) ys" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 479 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 480 | code_pred (expected_modes: (i => bool) => i => i => bool, (i => bool) => i => o => bool) filter2 . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 481 | code_pred [dseq] filter2 . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 482 | code_pred [random_dseq] filter2 . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 483 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 484 | thm filter2.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 485 | thm filter2.random_dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 486 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 487 | (* | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 488 | inductive filter3 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 489 | for P | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 490 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 491 | "List.filter P xs = ys ==> filter3 P xs ys" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 492 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 493 | code_pred (expected_modes: (o => bool) => i => o => bool, (o => bool) => i => i => bool , (i => bool) => i => o => bool, (i => bool) => i => i => bool) [skip_proof] filter3 . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 494 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 495 | code_pred [dseq] filter3 . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 496 | thm filter3.dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 497 | *) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 498 | (* | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 499 | inductive filter4 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 500 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 501 | "List.filter P xs = ys ==> filter4 P xs ys" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 502 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 503 | code_pred (expected_modes: i => i => o => bool, i => i => i => bool) filter4 . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 504 | (*code_pred [depth_limited] filter4 .*) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 505 | (*code_pred [random] filter4 .*) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 506 | *) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 507 | subsection {* reverse predicate *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 508 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 509 | inductive rev where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 510 | "rev [] []" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 511 | | "rev xs xs' ==> append xs' [x] ys ==> rev (x#xs) ys" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 512 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 513 | code_pred (expected_modes: i => o => bool, o => i => bool, i => i => bool) rev . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 514 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 515 | thm rev.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 516 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 517 | values "{xs. rev [0, 1, 2, 3::nat] xs}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 518 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 519 | inductive tupled_rev where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 520 | "tupled_rev ([], [])" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 521 | | "tupled_rev (xs, xs') \<Longrightarrow> tupled_append (xs', [x], ys) \<Longrightarrow> tupled_rev (x#xs, ys)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 522 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 523 | code_pred (expected_modes: i * o => bool, o * i => bool, i * i => bool) tupled_rev . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 524 | thm tupled_rev.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 525 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 526 | subsection {* partition predicate *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 527 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 528 | inductive partition :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 529 | for f where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 530 | "partition f [] [] []" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 531 | | "f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) (x # ys) zs" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 532 | | "\<not> f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) ys (x # zs)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 533 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 534 | code_pred (expected_modes: (i => bool) => i => o => o => bool, (i => bool) => o => i => i => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 535 | (i => bool) => i => i => o => bool, (i => bool) => i => o => i => bool, (i => bool) => i => i => i => bool) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 536 | partition . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 537 | code_pred [dseq] partition . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 538 | code_pred [random_dseq] partition . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 539 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 540 | values 10 "{(ys, zs). partition is_even
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 541 | [0, Suc 0, 2, 3, 4, 5, 6, 7] ys zs}" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 542 | values 10 "{zs. partition is_even zs [0, 2] [3, 5]}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 543 | values 10 "{zs. partition is_even zs [0, 7] [3, 5]}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 544 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 545 | inductive tupled_partition :: "('a \<Rightarrow> bool) \<Rightarrow> ('a list \<times> 'a list \<times> 'a list) \<Rightarrow> bool"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 546 | for f where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 547 | "tupled_partition f ([], [], [])" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 548 | | "f x \<Longrightarrow> tupled_partition f (xs, ys, zs) \<Longrightarrow> tupled_partition f (x # xs, x # ys, zs)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 549 | | "\<not> f x \<Longrightarrow> tupled_partition f (xs, ys, zs) \<Longrightarrow> tupled_partition f (x # xs, ys, x # zs)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 550 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 551 | code_pred (expected_modes: (i => bool) => i => bool, (i => bool) => (i * i * o) => bool, (i => bool) => (i * o * i) => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 552 | (i => bool) => (o * i * i) => bool, (i => bool) => (i * o * o) => bool) tupled_partition . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 553 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 554 | thm tupled_partition.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 555 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 556 | lemma [code_pred_intro]: | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 557 | "r a b \<Longrightarrow> tranclp r a b" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 558 | "r a b \<Longrightarrow> tranclp r b c \<Longrightarrow> tranclp r a c" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 559 | by auto | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 560 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 561 | subsection {* transitive predicate *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 562 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 563 | text {* Also look at the tabled transitive closure in the Library *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 564 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 565 | code_pred (modes: (i => o => bool) => i => i => bool, (i => o => bool) => i => o => bool as forwards_trancl, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 566 | (o => i => bool) => i => i => bool, (o => i => bool) => o => i => bool as backwards_trancl, (o => o => bool) => i => i => bool, (o => o => bool) => i => o => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 567 | (o => o => bool) => o => i => bool, (o => o => bool) => o => o => bool) tranclp | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 568 | proof - | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 569 | case tranclp | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 570 | from this converse_tranclpE[OF this(1)] show thesis by metis | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 571 | qed | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 572 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 573 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 574 | code_pred [dseq] tranclp . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 575 | code_pred [random_dseq] tranclp . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 576 | thm tranclp.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 577 | thm tranclp.random_dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 578 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 579 | inductive rtrancl' :: "'a => 'a => ('a => 'a => bool) => bool" 
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 580 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 581 | "rtrancl' x x r" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 582 | | "r x y ==> rtrancl' y z r ==> rtrancl' x z r" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 583 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 584 | code_pred [random_dseq] rtrancl' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 585 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 586 | thm rtrancl'.random_dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 587 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 588 | inductive rtrancl'' :: "('a * 'a * ('a \<Rightarrow> 'a \<Rightarrow> bool)) \<Rightarrow> bool"  
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 589 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 590 | "rtrancl'' (x, x, r)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 591 | | "r x y \<Longrightarrow> rtrancl'' (y, z, r) \<Longrightarrow> rtrancl'' (x, z, r)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 592 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 593 | code_pred rtrancl'' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 594 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 595 | inductive rtrancl''' :: "('a * ('a * 'a) * ('a * 'a => bool)) => bool" 
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 596 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 597 | "rtrancl''' (x, (x, x), r)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 598 | | "r (x, y) ==> rtrancl''' (y, (z, z), r) ==> rtrancl''' (x, (z, z), r)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 599 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 600 | code_pred rtrancl''' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 601 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 602 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 603 | inductive succ :: "nat \<Rightarrow> nat \<Rightarrow> bool" where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 604 | "succ 0 1" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 605 | | "succ m n \<Longrightarrow> succ (Suc m) (Suc n)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 606 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 607 | code_pred (modes: i => i => bool, i => o => bool, o => i => bool, o => o => bool) succ . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 608 | code_pred [random_dseq] succ . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 609 | thm succ.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 610 | thm succ.random_dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 611 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 612 | values 10 "{(m, n). succ n m}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 613 | values "{m. succ 0 m}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 614 | values "{m. succ m 0}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 615 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 616 | text {* values command needs mode annotation of the parameter succ
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 617 | to disambiguate which mode is to be chosen. *} | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 618 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 619 | values [mode: i => o => bool] 20 "{n. tranclp succ 10 n}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 620 | values [mode: o => i => bool] 10 "{n. tranclp succ n 10}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 621 | values 20 "{(n, m). tranclp succ n m}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 622 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 623 | inductive example_graph :: "int => int => bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 624 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 625 | "example_graph 0 1" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 626 | | "example_graph 1 2" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 627 | | "example_graph 1 3" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 628 | | "example_graph 4 7" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 629 | | "example_graph 4 5" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 630 | | "example_graph 5 6" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 631 | | "example_graph 7 6" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 632 | | "example_graph 7 8" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 633 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 634 | inductive not_reachable_in_example_graph :: "int => int => bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 635 | where "\<not> (tranclp example_graph x y) ==> not_reachable_in_example_graph x y" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 636 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 637 | code_pred (expected_modes: i => i => bool) not_reachable_in_example_graph . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 638 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 639 | thm not_reachable_in_example_graph.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 640 | thm tranclp.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 641 | value "not_reachable_in_example_graph 0 3" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 642 | value "not_reachable_in_example_graph 4 8" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 643 | value "not_reachable_in_example_graph 5 6" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 644 | text {* rtrancl compilation is strange! *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 645 | (* | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 646 | value "not_reachable_in_example_graph 0 4" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 647 | value "not_reachable_in_example_graph 1 6" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 648 | value "not_reachable_in_example_graph 8 4"*) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 649 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 650 | code_pred [dseq] not_reachable_in_example_graph . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 651 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 652 | values [dseq 6] "{x. tranclp example_graph 0 3}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 653 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 654 | values [dseq 0] "{x. not_reachable_in_example_graph 0 3}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 655 | values [dseq 0] "{x. not_reachable_in_example_graph 0 4}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 656 | values [dseq 20] "{x. not_reachable_in_example_graph 0 4}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 657 | values [dseq 6] "{x. not_reachable_in_example_graph 0 3}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 658 | values [dseq 3] "{x. not_reachable_in_example_graph 4 2}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 659 | values [dseq 6] "{x. not_reachable_in_example_graph 4 2}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 660 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 661 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 662 | inductive not_reachable_in_example_graph' :: "int => int => bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 663 | where "\<not> (rtranclp example_graph x y) ==> not_reachable_in_example_graph' x y" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 664 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 665 | code_pred not_reachable_in_example_graph' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 666 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 667 | value "not_reachable_in_example_graph' 0 3" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 668 | (* value "not_reachable_in_example_graph' 0 5" would not terminate *) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 669 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 670 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 671 | (*values [depth_limited 0] "{x. not_reachable_in_example_graph' 0 3}"*)
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 672 | (*values [depth_limited 3] "{x. not_reachable_in_example_graph' 0 3}"*) (* fails with undefined *)
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 673 | (*values [depth_limited 5] "{x. not_reachable_in_example_graph' 0 3}"*)
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 674 | (*values [depth_limited 1] "{x. not_reachable_in_example_graph' 0 4}"*)
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 675 | (*values [depth_limit = 4] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 676 | (*values [depth_limit = 20] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 677 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 678 | code_pred [dseq] not_reachable_in_example_graph' . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 679 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 680 | (*thm not_reachable_in_example_graph'.dseq_equation*) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 681 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 682 | (*values [dseq 0] "{x. not_reachable_in_example_graph' 0 3}"*)
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 683 | (*values [depth_limited 3] "{x. not_reachable_in_example_graph' 0 3}"*) (* fails with undefined *)
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 684 | (*values [depth_limited 5] "{x. not_reachable_in_example_graph' 0 3}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 685 | values [depth_limited 1] "{x. not_reachable_in_example_graph' 0 4}"*)
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 686 | (*values [depth_limit = 4] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 687 | (*values [depth_limit = 20] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 688 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 689 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 690 | subsection {* IMP *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 691 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 692 | types | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 693 | var = nat | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 694 | state = "int list" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 695 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 696 | datatype com = | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 697 | Skip | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 698 | Ass var "state => int" | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 699 | Seq com com | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 700 | IF "state => bool" com com | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 701 | While "state => bool" com | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 702 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 703 | inductive exec :: "com => state => state => bool" where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 704 | "exec Skip s s" | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 705 | "exec (Ass x e) s (s[x := e(s)])" | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 706 | "exec c1 s1 s2 ==> exec c2 s2 s3 ==> exec (Seq c1 c2) s1 s3" | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 707 | "b s ==> exec c1 s t ==> exec (IF b c1 c2) s t" | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 708 | "~b s ==> exec c2 s t ==> exec (IF b c1 c2) s t" | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 709 | "~b s ==> exec (While b c) s s" | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 710 | "b s1 ==> exec c s1 s2 ==> exec (While b c) s2 s3 ==> exec (While b c) s1 s3" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 711 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 712 | code_pred exec . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 713 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 714 | values "{t. exec
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 715 | (While (%s. s!0 > 0) (Seq (Ass 0 (%s. s!0 - 1)) (Ass 1 (%s. s!1 + 1)))) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 716 | [3,5] t}" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 717 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 718 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 719 | inductive tupled_exec :: "(com \<times> state \<times> state) \<Rightarrow> bool" where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 720 | "tupled_exec (Skip, s, s)" | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 721 | "tupled_exec (Ass x e, s, s[x := e(s)])" | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 722 | "tupled_exec (c1, s1, s2) ==> tupled_exec (c2, s2, s3) ==> tupled_exec (Seq c1 c2, s1, s3)" | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 723 | "b s ==> tupled_exec (c1, s, t) ==> tupled_exec (IF b c1 c2, s, t)" | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 724 | "~b s ==> tupled_exec (c2, s, t) ==> tupled_exec (IF b c1 c2, s, t)" | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 725 | "~b s ==> tupled_exec (While b c, s, s)" | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 726 | "b s1 ==> tupled_exec (c, s1, s2) ==> tupled_exec (While b c, s2, s3) ==> tupled_exec (While b c, s1, s3)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 727 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 728 | code_pred tupled_exec . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 729 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 730 | values "{s. tupled_exec (While (%s. s!0 > 0) (Seq (Ass 0 (%s. s!0 - 1)) (Ass 1 (%s. s!1 + 1))), [3, 5], s)}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 731 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 732 | subsection {* CCS *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 733 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 734 | text{* This example formalizes finite CCS processes without communication or
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 735 | recursion. For simplicity, labels are natural numbers. *} | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 736 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 737 | datatype proc = nil | pre nat proc | or proc proc | par proc proc | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 738 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 739 | inductive step :: "proc \<Rightarrow> nat \<Rightarrow> proc \<Rightarrow> bool" where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 740 | "step (pre n p) n p" | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 741 | "step p1 a q \<Longrightarrow> step (or p1 p2) a q" | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 742 | "step p2 a q \<Longrightarrow> step (or p1 p2) a q" | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 743 | "step p1 a q \<Longrightarrow> step (par p1 p2) a (par q p2)" | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 744 | "step p2 a q \<Longrightarrow> step (par p1 p2) a (par p1 q)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 745 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 746 | code_pred step . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 747 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 748 | inductive steps where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 749 | "steps p [] p" | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 750 | "step p a q \<Longrightarrow> steps q as r \<Longrightarrow> steps p (a#as) r" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 751 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 752 | code_pred steps . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 753 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 754 | values 3 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 755 |  "{as . steps (par (or (pre 0 nil) (pre 1 nil)) (pre 2 nil)) as (par nil nil)}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 756 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 757 | values 5 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 758 |  "{as . steps (par (or (pre 0 nil) (pre 1 nil)) (pre 2 nil)) as (par nil nil)}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 759 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 760 | values 3 "{(a,q). step (par nil nil) a q}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 761 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 762 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 763 | inductive tupled_step :: "(proc \<times> nat \<times> proc) \<Rightarrow> bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 764 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 765 | "tupled_step (pre n p, n, p)" | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 766 | "tupled_step (p1, a, q) \<Longrightarrow> tupled_step (or p1 p2, a, q)" | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 767 | "tupled_step (p2, a, q) \<Longrightarrow> tupled_step (or p1 p2, a, q)" | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 768 | "tupled_step (p1, a, q) \<Longrightarrow> tupled_step (par p1 p2, a, par q p2)" | | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 769 | "tupled_step (p2, a, q) \<Longrightarrow> tupled_step (par p1 p2, a, par p1 q)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 770 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 771 | code_pred tupled_step . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 772 | thm tupled_step.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 773 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 774 | subsection {* divmod *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 775 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 776 | inductive divmod_rel :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 777 | "k < l \<Longrightarrow> divmod_rel k l 0 k" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 778 | | "k \<ge> l \<Longrightarrow> divmod_rel (k - l) l q r \<Longrightarrow> divmod_rel k l (Suc q) r" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 779 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 780 | code_pred divmod_rel .. | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 781 | thm divmod_rel.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 782 | value [code] "Predicate.the (divmod_rel_i_i_o_o 1705 42)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 783 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 784 | subsection {* Transforming predicate logic into logic programs *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 785 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 786 | subsection {* Transforming functions into logic programs *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 787 | definition | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 788 | "case_f xs ys = (case (xs @ ys) of [] => [] | (x # xs) => xs)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 789 | |
| 36040 
fcd7bea01a93
adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
 bulwahn parents: 
35954diff
changeset | 790 | code_pred [inductify, skip_proof] case_f . | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 791 | thm case_fP.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 792 | thm case_fP.intros | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 793 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 794 | fun fold_map_idx where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 795 | "fold_map_idx f i y [] = (y, [])" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 796 | | "fold_map_idx f i y (x # xs) = | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 797 | (let (y', x') = f i y x; (y'', xs') = fold_map_idx f (Suc i) y' xs | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 798 | in (y'', x' # xs'))" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 799 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 800 | text {* mode analysis explores thousand modes - this is infeasible at the moment... *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 801 | (*code_pred [inductify, show_steps] fold_map_idx .*) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 802 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 803 | subsection {* Minimum *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 804 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 805 | definition Min | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 806 | where "Min s r x \<equiv> s x \<and> (\<forall>y. r x y \<longrightarrow> x = y)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 807 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 808 | code_pred [inductify] Min . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 809 | thm Min.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 810 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 811 | subsection {* Lexicographic order *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 812 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 813 | declare lexord_def[code_pred_def] | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 814 | code_pred [inductify] lexord . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 815 | code_pred [random_dseq inductify] lexord . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 816 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 817 | thm lexord.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 818 | thm lexord.random_dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 819 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 820 | inductive less_than_nat :: "nat * nat => bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 821 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 822 | "less_than_nat (0, x)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 823 | | "less_than_nat (x, y) ==> less_than_nat (Suc x, Suc y)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 824 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 825 | code_pred less_than_nat . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 826 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 827 | code_pred [dseq] less_than_nat . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 828 | code_pred [random_dseq] less_than_nat . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 829 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 830 | inductive test_lexord :: "nat list * nat list => bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 831 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 832 | "lexord less_than_nat (xs, ys) ==> test_lexord (xs, ys)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 833 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 834 | code_pred test_lexord . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 835 | code_pred [dseq] test_lexord . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 836 | code_pred [random_dseq] test_lexord . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 837 | thm test_lexord.dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 838 | thm test_lexord.random_dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 839 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 840 | values "{x. test_lexord ([1, 2, 3], [1, 2, 5])}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 841 | (*values [depth_limited 5] "{x. test_lexord ([1, 2, 3], [1, 2, 5])}"*)
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 842 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 843 | lemmas [code_pred_def] = lexn_conv lex_conv lenlex_conv | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 844 | (* | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 845 | code_pred [inductify] lexn . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 846 | thm lexn.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 847 | *) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 848 | (* | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 849 | code_pred [random_dseq inductify] lexn . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 850 | thm lexn.random_dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 851 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 852 | values [random_dseq 4, 4, 6] 100 "{(n, xs, ys::int list). lexn (%(x, y). x <= y) n (xs, ys)}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 853 | *) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 854 | inductive has_length | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 855 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 856 | "has_length [] 0" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 857 | | "has_length xs i ==> has_length (x # xs) (Suc i)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 858 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 859 | lemma has_length: | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 860 | "has_length xs n = (length xs = n)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 861 | proof (rule iffI) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 862 | assume "has_length xs n" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 863 | from this show "length xs = n" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 864 | by (rule has_length.induct) auto | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 865 | next | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 866 | assume "length xs = n" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 867 | from this show "has_length xs n" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 868 | by (induct xs arbitrary: n) (auto intro: has_length.intros) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 869 | qed | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 870 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 871 | lemma lexn_intros [code_pred_intro]: | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 872 | "has_length xs i ==> has_length ys i ==> r (x, y) ==> lexn r (Suc i) (x # xs, y # ys)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 873 | "lexn r i (xs, ys) ==> lexn r (Suc i) (x # xs, x # ys)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 874 | proof - | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 875 | assume "has_length xs i" "has_length ys i" "r (x, y)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 876 | from this has_length show "lexn r (Suc i) (x # xs, y # ys)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 877 | unfolding lexn_conv Collect_def mem_def | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 878 | by fastsimp | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 879 | next | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 880 | assume "lexn r i (xs, ys)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 881 | thm lexn_conv | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 882 | from this show "lexn r (Suc i) (x#xs, x#ys)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 883 | unfolding Collect_def mem_def lexn_conv | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 884 | apply auto | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 885 | apply (rule_tac x="x # xys" in exI) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 886 | by auto | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 887 | qed | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 888 | |
| 36055 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 889 | code_pred [random_dseq] lexn | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 890 | proof - | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 891 | fix r n xs ys | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 892 | assume 1: "lexn r n (xs, ys)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 893 | assume 2: "\<And>r' i x xs' y ys'. r = r' ==> n = Suc i ==> (xs, ys) = (x # xs', y # ys') ==> has_length xs' i ==> has_length ys' i ==> r' (x, y) ==> thesis" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 894 | assume 3: "\<And>r' i x xs' ys'. r = r' ==> n = Suc i ==> (xs, ys) = (x # xs', x # ys') ==> lexn r' i (xs', ys') ==> thesis" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 895 | from 1 2 3 show thesis | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 896 | unfolding lexn_conv Collect_def mem_def | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 897 | apply (auto simp add: has_length) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 898 | apply (case_tac xys) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 899 | apply auto | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 900 | apply fastsimp | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 901 | apply fastsimp done | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 902 | qed | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 903 | |
| 36055 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 904 | values [random_dseq 1, 2, 5] 10 "{(n, xs, ys::int list). lexn (%(x, y). x <= y) n (xs, ys)}"
 | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 905 | |
| 36040 
fcd7bea01a93
adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
 bulwahn parents: 
35954diff
changeset | 906 | code_pred [inductify, skip_proof] lex . | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 907 | thm lex.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 908 | thm lex_def | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 909 | declare lenlex_conv[code_pred_def] | 
| 36040 
fcd7bea01a93
adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
 bulwahn parents: 
35954diff
changeset | 910 | code_pred [inductify, skip_proof] lenlex . | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 911 | thm lenlex.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 912 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 913 | code_pred [random_dseq inductify] lenlex . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 914 | thm lenlex.random_dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 915 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 916 | values [random_dseq 4, 2, 4] 100 "{(xs, ys::int list). lenlex (%(x, y). x <= y) (xs, ys)}"
 | 
| 36055 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 917 | |
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 918 | thm lists.intros | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 919 | code_pred [inductify] lists . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 920 | thm lists.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 921 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 922 | subsection {* AVL Tree *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 923 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 924 | datatype 'a tree = ET | MKT 'a "'a tree" "'a tree" nat | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 925 | fun height :: "'a tree => nat" where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 926 | "height ET = 0" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 927 | | "height (MKT x l r h) = max (height l) (height r) + 1" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 928 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 929 | consts avl :: "'a tree => bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 930 | primrec | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 931 | "avl ET = True" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 932 | "avl (MKT x l r h) = ((height l = height r \<or> height l = 1 + height r \<or> height r = 1+height l) \<and> | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 933 | h = max (height l) (height r) + 1 \<and> avl l \<and> avl r)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 934 | (* | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 935 | code_pred [inductify] avl . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 936 | thm avl.equation*) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 937 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 938 | code_pred [random_dseq inductify] avl . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 939 | thm avl.random_dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 940 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 941 | values [random_dseq 2, 1, 7] 5 "{t:: int tree. avl t}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 942 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 943 | fun set_of | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 944 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 945 | "set_of ET = {}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 946 | | "set_of (MKT n l r h) = insert n (set_of l \<union> set_of r)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 947 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 948 | fun is_ord :: "nat tree => bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 949 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 950 | "is_ord ET = True" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 951 | | "is_ord (MKT n l r h) = | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 952 | ((\<forall>n' \<in> set_of l. n' < n) \<and> (\<forall>n' \<in> set_of r. n < n') \<and> is_ord l \<and> is_ord r)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 953 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 954 | code_pred (expected_modes: i => o => bool, i => i => bool) [inductify] set_of . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 955 | thm set_of.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 956 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 957 | code_pred (expected_modes: i => bool) [inductify] is_ord . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 958 | thm is_ord_aux.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 959 | thm is_ord.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 960 | |
| 36055 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 961 | subsection {* Definitions about Relations *}
 | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 962 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 963 | term "converse" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 964 | code_pred (modes: | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 965 | (i * i => bool) => i * i => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 966 | (i * o => bool) => o * i => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 967 | (i * o => bool) => i * i => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 968 | (o * i => bool) => i * o => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 969 | (o * i => bool) => i * i => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 970 | (o * o => bool) => o * o => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 971 | (o * o => bool) => i * o => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 972 | (o * o => bool) => o * i => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 973 | (o * o => bool) => i * i => bool) [inductify] converse . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 974 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 975 | thm converse.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 976 | code_pred [inductify] rel_comp . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 977 | thm rel_comp.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 978 | code_pred [inductify] Image . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 979 | thm Image.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 980 | declare singleton_iff[code_pred_inline] | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 981 | declare Id_on_def[unfolded Bex_def UNION_def singleton_iff, code_pred_def] | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 982 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 983 | code_pred (expected_modes: | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 984 | (o => bool) => o => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 985 | (o => bool) => i * o => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 986 | (o => bool) => o * i => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 987 | (o => bool) => i => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 988 | (i => bool) => i * o => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 989 | (i => bool) => o * i => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 990 | (i => bool) => i => bool) [inductify] Id_on . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 991 | thm Id_on.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 992 | thm Domain_def | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 993 | code_pred (modes: | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 994 | (o * o => bool) => o => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 995 | (o * o => bool) => i => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 996 | (i * o => bool) => i => bool) [inductify] Domain . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 997 | thm Domain.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 998 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 999 | thm Range_def | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1000 | code_pred (modes: | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1001 | (o * o => bool) => o => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1002 | (o * o => bool) => i => bool, | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1003 | (o * i => bool) => i => bool) [inductify] Range . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1004 | thm Range.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1005 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1006 | code_pred [inductify] Field . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1007 | thm Field.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1008 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1009 | thm refl_on_def | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1010 | code_pred [inductify] refl_on . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1011 | thm refl_on.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1012 | code_pred [inductify] total_on . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1013 | thm total_on.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1014 | code_pred [inductify] antisym . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1015 | thm antisym.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1016 | code_pred [inductify] trans . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1017 | thm trans.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1018 | code_pred [inductify] single_valued . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1019 | thm single_valued.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1020 | thm inv_image_def | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1021 | code_pred [inductify] inv_image . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1022 | thm inv_image.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1023 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1024 | subsection {* Inverting list functions *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1025 | |
| 36055 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1026 | code_pred [inductify] size_list . | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1027 | code_pred [new_random_dseq inductify] size_list . | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1028 | thm size_listP.equation | 
| 36055 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1029 | thm size_listP.new_random_dseq_equation | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1030 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1031 | values [new_random_dseq 2,3,10] 3 "{xs. size_listP (xs::nat list) (5::nat)}"
 | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1032 | |
| 36040 
fcd7bea01a93
adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
 bulwahn parents: 
35954diff
changeset | 1033 | code_pred (expected_modes: i => o => bool, o => i => bool, i => i => bool) [inductify, skip_proof] List.concat . | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1034 | thm concatP.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1035 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1036 | values "{ys. concatP [[1, 2], [3, (4::int)]] ys}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1037 | values "{ys. concatP [[1, 2], [3]] [1, 2, (3::nat)]}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1038 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1039 | code_pred [dseq inductify] List.concat . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1040 | thm concatP.dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1041 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1042 | values [dseq 3] 3 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1043 |   "{xs. concatP xs ([0] :: nat list)}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1044 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1045 | values [dseq 5] 3 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1046 |   "{xs. concatP xs ([1] :: int list)}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1047 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1048 | values [dseq 5] 3 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1049 |   "{xs. concatP xs ([1] :: nat list)}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1050 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1051 | values [dseq 5] 3 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1052 |   "{xs. concatP xs [(1::int), 2]}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1053 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1054 | code_pred (expected_modes: i => o => bool, i => i => bool) [inductify] hd . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1055 | thm hdP.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1056 | values "{x. hdP [1, 2, (3::int)] x}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1057 | values "{(xs, x). hdP [1, 2, (3::int)] 1}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1058 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1059 | code_pred (expected_modes: i => o => bool, i => i => bool) [inductify] tl . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1060 | thm tlP.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1061 | values "{x. tlP [1, 2, (3::nat)] x}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1062 | values "{x. tlP [1, 2, (3::int)] [3]}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1063 | |
| 36040 
fcd7bea01a93
adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
 bulwahn parents: 
35954diff
changeset | 1064 | code_pred [inductify, skip_proof] last . | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1065 | thm lastP.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1066 | |
| 36040 
fcd7bea01a93
adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
 bulwahn parents: 
35954diff
changeset | 1067 | code_pred [inductify, skip_proof] butlast . | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1068 | thm butlastP.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1069 | |
| 36040 
fcd7bea01a93
adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
 bulwahn parents: 
35954diff
changeset | 1070 | code_pred [inductify, skip_proof] take . | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1071 | thm takeP.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1072 | |
| 36040 
fcd7bea01a93
adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
 bulwahn parents: 
35954diff
changeset | 1073 | code_pred [inductify, skip_proof] drop . | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1074 | thm dropP.equation | 
| 36040 
fcd7bea01a93
adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
 bulwahn parents: 
35954diff
changeset | 1075 | code_pred [inductify, skip_proof] zip . | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1076 | thm zipP.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1077 | |
| 36040 
fcd7bea01a93
adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
 bulwahn parents: 
35954diff
changeset | 1078 | code_pred [inductify, skip_proof] upt . | 
| 
fcd7bea01a93
adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
 bulwahn parents: 
35954diff
changeset | 1079 | code_pred [inductify, skip_proof] remdups . | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1080 | thm remdupsP.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1081 | code_pred [dseq inductify] remdups . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1082 | values [dseq 4] 5 "{xs. remdupsP xs [1, (2::int)]}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1083 | |
| 36040 
fcd7bea01a93
adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
 bulwahn parents: 
35954diff
changeset | 1084 | code_pred [inductify, skip_proof] remove1 . | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1085 | thm remove1P.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1086 | values "{xs. remove1P 1 xs [2, (3::int)]}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1087 | |
| 36040 
fcd7bea01a93
adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
 bulwahn parents: 
35954diff
changeset | 1088 | code_pred [inductify, skip_proof] removeAll . | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1089 | thm removeAllP.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1090 | code_pred [dseq inductify] removeAll . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1091 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1092 | values [dseq 4] 10 "{xs. removeAllP 1 xs [(2::nat)]}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1093 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1094 | code_pred [inductify] distinct . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1095 | thm distinct.equation | 
| 36040 
fcd7bea01a93
adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
 bulwahn parents: 
35954diff
changeset | 1096 | code_pred [inductify, skip_proof] replicate . | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1097 | thm replicateP.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1098 | values 5 "{(n, xs). replicateP n (0::int) xs}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1099 | |
| 36040 
fcd7bea01a93
adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
 bulwahn parents: 
35954diff
changeset | 1100 | code_pred [inductify, skip_proof] splice . | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1101 | thm splice.simps | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1102 | thm spliceP.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1103 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1104 | values "{xs. spliceP xs [1, 2, 3] [1, 1, 1, 2, 1, (3::nat)]}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1105 | |
| 36040 
fcd7bea01a93
adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
 bulwahn parents: 
35954diff
changeset | 1106 | code_pred [inductify, skip_proof] List.rev . | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1107 | code_pred [inductify] map . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1108 | code_pred [inductify] foldr . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1109 | code_pred [inductify] foldl . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1110 | code_pred [inductify] filter . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1111 | code_pred [random_dseq inductify] filter . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1112 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1113 | subsection {* Context Free Grammar *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1114 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1115 | datatype alphabet = a | b | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1116 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1117 | inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1118 | "[] \<in> S\<^isub>1" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1119 | | "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1120 | | "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1121 | | "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1122 | | "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1123 | | "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1124 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1125 | code_pred [inductify] S\<^isub>1p . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1126 | code_pred [random_dseq inductify] S\<^isub>1p . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1127 | thm S\<^isub>1p.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1128 | thm S\<^isub>1p.random_dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1129 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1130 | values [random_dseq 5, 5, 5] 5 "{x. S\<^isub>1p x}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1131 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1132 | inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1133 | "[] \<in> S\<^isub>2" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1134 | | "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1135 | | "w \<in> B\<^isub>2 \<Longrightarrow> a # w \<in> S\<^isub>2" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1136 | | "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1137 | | "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1138 | | "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1139 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1140 | code_pred [random_dseq inductify] S\<^isub>2p . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1141 | thm S\<^isub>2p.random_dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1142 | thm A\<^isub>2p.random_dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1143 | thm B\<^isub>2p.random_dseq_equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1144 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1145 | values [random_dseq 5, 5, 5] 10 "{x. S\<^isub>2p x}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1146 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1147 | inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1148 | "[] \<in> S\<^isub>3" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1149 | | "w \<in> A\<^isub>3 \<Longrightarrow> b # w \<in> S\<^isub>3" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1150 | | "w \<in> B\<^isub>3 \<Longrightarrow> a # w \<in> S\<^isub>3" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1151 | | "w \<in> S\<^isub>3 \<Longrightarrow> a # w \<in> A\<^isub>3" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1152 | | "w \<in> S\<^isub>3 \<Longrightarrow> b # w \<in> B\<^isub>3" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1153 | | "\<lbrakk>v \<in> B\<^isub>3; w \<in> B\<^isub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>3" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1154 | |
| 36040 
fcd7bea01a93
adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
 bulwahn parents: 
35954diff
changeset | 1155 | code_pred [inductify, skip_proof] S\<^isub>3p . | 
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1156 | thm S\<^isub>3p.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1157 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1158 | values 10 "{x. S\<^isub>3p x}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1159 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1160 | inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1161 | "[] \<in> S\<^isub>4" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1162 | | "w \<in> A\<^isub>4 \<Longrightarrow> b # w \<in> S\<^isub>4" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1163 | | "w \<in> B\<^isub>4 \<Longrightarrow> a # w \<in> S\<^isub>4" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1164 | | "w \<in> S\<^isub>4 \<Longrightarrow> a # w \<in> A\<^isub>4" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1165 | | "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1166 | | "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1167 | | "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1168 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1169 | code_pred (expected_modes: o => bool, i => bool) S\<^isub>4p . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1170 | |
| 36055 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1171 | hide const a b | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1172 | |
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1173 | subsection {* Lambda *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1174 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1175 | datatype type = | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1176 | Atom nat | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1177 | | Fun type type (infixr "\<Rightarrow>" 200) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1178 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1179 | datatype dB = | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1180 | Var nat | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1181 | | App dB dB (infixl "\<degree>" 200) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1182 | | Abs type dB | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1183 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1184 | primrec | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1185 |   nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>" [90, 0] 91)
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1186 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1187 | "[]\<langle>i\<rangle> = None" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1188 | | "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1189 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1190 | inductive nth_el' :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1191 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1192 | "nth_el' (x # xs) 0 x" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1193 | | "nth_el' xs i y \<Longrightarrow> nth_el' (x # xs) (Suc i) y" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1194 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1195 | inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool"  ("_ \<turnstile> _ : _" [50, 50, 50] 50)
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1196 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1197 | Var [intro!]: "nth_el' env x T \<Longrightarrow> env \<turnstile> Var x : T" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1198 | | Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> U)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1199 | | App [intro!]: "env \<turnstile> s : T \<Rightarrow> U \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1200 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1201 | primrec | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1202 | lift :: "[dB, nat] => dB" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1203 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1204 | "lift (Var i) k = (if i < k then Var i else Var (i + 1))" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1205 | | "lift (s \<degree> t) k = lift s k \<degree> lift t k" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1206 | | "lift (Abs T s) k = Abs T (lift s (k + 1))" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1207 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1208 | primrec | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1209 |   subst :: "[dB, dB, nat] => dB"  ("_[_'/_]" [300, 0, 0] 300)
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1210 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1211 | subst_Var: "(Var i)[s/k] = | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1212 | (if k < i then Var (i - 1) else if i = k then s else Var i)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1213 | | subst_App: "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1214 | | subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1215 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1216 | inductive beta :: "[dB, dB] => bool" (infixl "\<rightarrow>\<^sub>\<beta>" 50) | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1217 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1218 | beta [simp, intro!]: "Abs T s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1219 | | appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1220 | | appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1221 | | abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1222 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1223 | code_pred (expected_modes: i => i => o => bool, i => i => i => bool) typing . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1224 | thm typing.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1225 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1226 | code_pred (modes: i => i => bool, i => o => bool as reduce') beta . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1227 | thm beta.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1228 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1229 | values "{x. App (Abs (Atom 0) (Var 0)) (Var 1) \<rightarrow>\<^sub>\<beta> x}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1230 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1231 | definition "reduce t = Predicate.the (reduce' t)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1232 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1233 | value "reduce (App (Abs (Atom 0) (Var 0)) (Var 1))" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1234 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1235 | code_pred [dseq] typing . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1236 | code_pred [random_dseq] typing . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1237 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1238 | values [random_dseq 1,1,5] 10 "{(\<Gamma>, t, T). \<Gamma> \<turnstile> t : T}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1239 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1240 | subsection {* A minimal example of yet another semantics *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1241 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1242 | text {* thanks to Elke Salecker *}
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1243 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1244 | types | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1245 | vname = nat | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1246 | vvalue = int | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1247 | var_assign = "vname \<Rightarrow> vvalue" --"variable assignment" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1248 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1249 | datatype ir_expr = | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1250 | IrConst vvalue | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1251 | | ObjAddr vname | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1252 | | Add ir_expr ir_expr | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1253 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1254 | datatype val = | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1255 | IntVal vvalue | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1256 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1257 | record configuration = | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1258 | Env :: var_assign | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1259 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1260 | inductive eval_var :: | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1261 | "ir_expr \<Rightarrow> configuration \<Rightarrow> val \<Rightarrow> bool" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1262 | where | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1263 | irconst: "eval_var (IrConst i) conf (IntVal i)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1264 | | objaddr: "\<lbrakk> Env conf n = i \<rbrakk> \<Longrightarrow> eval_var (ObjAddr n) conf (IntVal i)" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1265 | | plus: "\<lbrakk> eval_var l conf (IntVal vl); eval_var r conf (IntVal vr) \<rbrakk> \<Longrightarrow> eval_var (Add l r) conf (IntVal (vl+vr))" | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1266 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1267 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1268 | code_pred eval_var . | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1269 | thm eval_var.equation | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1270 | |
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1271 | values "{val. eval_var (Add (IrConst 1) (IrConst 2)) (| Env = (\<lambda>x. 0)|) val}"
 | 
| 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1272 | |
| 36055 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1273 | section {* Function predicate replacement *}
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1274 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1275 | text {*
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1276 | If the mode analysis uses the functional mode, we | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1277 | replace predicates that resulted from functions again by their functions. | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1278 | *} | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1279 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1280 | inductive test_append | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1281 | where | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1282 | "List.append xs ys = zs ==> test_append xs ys zs" | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1283 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1284 | code_pred [inductify, skip_proof] test_append . | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1285 | thm test_append.equation | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1286 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1287 | text {* If append is not turned into a predicate, then the mode
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1288 | o => o => i => bool could not be inferred. *} | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1289 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1290 | values 4 "{(xs::int list, ys). test_append xs ys [1, 2, 3, 4]}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1291 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1292 | text {* If appendP is not reverted back to a function, then mode i => i => o => bool
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1293 | fails after deleting the predicate equation. *} | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1294 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1295 | declare appendP.equation[code del] | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1296 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1297 | values "{xs::int list. test_append [1,2] [3,4] xs}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1298 | values "{xs::int list. test_append (replicate 1000 1) (replicate 1000 2) xs}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1299 | values "{xs::int list. test_append (replicate 2000 1) (replicate 2000 2) xs}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1300 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1301 | subsection {* Function with tuples *}
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1302 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1303 | fun append' | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1304 | where | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1305 | "append' ([], ys) = ys" | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1306 | | "append' (x # xs, ys) = x # append' (xs, ys)" | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1307 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1308 | inductive test_append' | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1309 | where | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1310 | "append' (xs, ys) = zs ==> test_append' xs ys zs" | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1311 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1312 | code_pred [inductify, skip_proof] test_append' . | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1313 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1314 | thm test_append'.equation | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1315 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1316 | values "{(xs::int list, ys). test_append' xs ys [1, 2, 3, 4]}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1317 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1318 | declare append'P.equation[code del] | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1319 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1320 | values "{zs :: int list. test_append' [1,2,3] [4,5] zs}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1321 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1322 | section {* Arithmetic examples *}
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1323 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1324 | subsection {* Examples on nat *}
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1325 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1326 | inductive plus_nat_test :: "nat => nat => nat => bool" | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1327 | where | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1328 | "x + y = z ==> plus_nat_test x y z" | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1329 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1330 | code_pred [inductify, skip_proof] plus_nat_test . | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1331 | code_pred [new_random_dseq inductify] plus_nat_test . | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1332 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1333 | thm plus_nat_test.equation | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1334 | thm plus_nat_test.new_random_dseq_equation | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1335 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1336 | values [expected "{9::nat}"] "{z. plus_nat_test 4 5 z}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1337 | values [expected "{9::nat}"] "{z. plus_nat_test 7 2 z}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1338 | values [expected "{4::nat}"] "{y. plus_nat_test 5 y 9}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1339 | values [expected "{}"] "{y. plus_nat_test 9 y 8}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1340 | values [expected "{6::nat}"] "{y. plus_nat_test 1 y 7}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1341 | values [expected "{2::nat}"] "{x. plus_nat_test x 7 9}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1342 | values [expected "{}"] "{x. plus_nat_test x 9 7}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1343 | values [expected "{(0::nat,0::nat)}"] "{(x, y). plus_nat_test x y 0}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1344 | values [expected "{(0, Suc 0), (Suc 0, 0)}"] "{(x, y). plus_nat_test x y 1}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1345 | values [expected "{(0, 5), (4, Suc 0), (3, 2), (2, 3), (Suc 0, 4), (5, 0)}"]
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1346 |   "{(x, y). plus_nat_test x y 5}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1347 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1348 | inductive minus_nat_test :: "nat => nat => nat => bool" | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1349 | where | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1350 | "x - y = z ==> minus_nat_test x y z" | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1351 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1352 | code_pred [inductify, skip_proof] minus_nat_test . | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1353 | code_pred [new_random_dseq inductify] minus_nat_test . | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1354 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1355 | thm minus_nat_test.equation | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1356 | thm minus_nat_test.new_random_dseq_equation | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1357 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1358 | values [expected "{0::nat}"] "{z. minus_nat_test 4 5 z}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1359 | values [expected "{5::nat}"] "{z. minus_nat_test 7 2 z}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1360 | values [expected "{16::nat}"] "{x. minus_nat_test x 7 9}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1361 | values [expected "{16::nat}"] "{x. minus_nat_test x 9 7}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1362 | values [expected "{0, Suc 0, 2, 3}"] "{x. minus_nat_test x 3 0}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1363 | values [expected "{0::nat}"] "{x. minus_nat_test x 0 0}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1364 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1365 | subsection {* Examples on int *}
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1366 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1367 | inductive plus_int_test :: "int => int => int => bool" | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1368 | where | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1369 | "a + b = c ==> plus_int_test a b c" | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1370 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1371 | code_pred [inductify, skip_proof] plus_int_test . | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1372 | code_pred [new_random_dseq inductify] plus_int_test . | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1373 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1374 | thm plus_int_test.equation | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1375 | thm plus_int_test.new_random_dseq_equation | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1376 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1377 | values [expected "{1::int}"] "{a. plus_int_test a 6 7}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1378 | values [expected "{1::int}"] "{b. plus_int_test 6 b 7}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1379 | values [expected "{11::int}"] "{c. plus_int_test 5 6 c}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1380 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1381 | inductive minus_int_test :: "int => int => int => bool" | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1382 | where | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1383 | "a - b = c ==> minus_int_test a b c" | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1384 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1385 | code_pred [inductify, skip_proof] minus_int_test . | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1386 | code_pred [new_random_dseq inductify] minus_int_test . | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1387 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1388 | thm minus_int_test.equation | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1389 | thm minus_int_test.new_random_dseq_equation | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1390 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1391 | values [expected "{4::int}"] "{c. minus_int_test 9 5 c}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1392 | values [expected "{9::int}"] "{a. minus_int_test a 4 5}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1393 | values [expected "{- 1::int}"] "{b. minus_int_test 4 b 5}"
 | 
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1394 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1395 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1396 | |
| 
537876d0fa62
adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
 bulwahn parents: 
36040diff
changeset | 1397 | |
| 35950 
791ce568d40a
moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
 bulwahn parents: diff
changeset | 1398 | end |