| author | wenzelm | 
| Mon, 16 Jan 2023 20:08:15 +0100 | |
| changeset 76994 | 7c23db6b857b | 
| parent 74365 | b49bd5d9041f | 
| permissions | -rw-r--r-- | 
| 
74365
 
b49bd5d9041f
improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
 
wenzelm 
parents: 
61343 
diff
changeset
 | 
1  | 
(* Title: Pure/ex/Guess_Examples.thy  | 
| 20005 | 2  | 
Author: Makarius  | 
3  | 
*)  | 
|
4  | 
||
| 
74365
 
b49bd5d9041f
improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
 
wenzelm 
parents: 
61343 
diff
changeset
 | 
5  | 
section \<open>Proof by wild guessing\<close>  | 
| 20005 | 6  | 
|
| 
74365
 
b49bd5d9041f
improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
 
wenzelm 
parents: 
61343 
diff
changeset
 | 
7  | 
theory Guess_Examples  | 
| 
 
b49bd5d9041f
improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
 
wenzelm 
parents: 
61343 
diff
changeset
 | 
8  | 
imports "Pure-Examples.Higher_Order_Logic" Guess  | 
| 20005 | 9  | 
begin  | 
10  | 
||
| 
74365
 
b49bd5d9041f
improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
 
wenzelm 
parents: 
61343 
diff
changeset
 | 
11  | 
typedecl t  | 
| 
 
b49bd5d9041f
improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
 
wenzelm 
parents: 
61343 
diff
changeset
 | 
12  | 
instance t :: type ..  | 
| 
 
b49bd5d9041f
improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
 
wenzelm 
parents: 
61343 
diff
changeset
 | 
13  | 
|
| 59031 | 14  | 
notepad  | 
15  | 
begin  | 
|
| 
74365
 
b49bd5d9041f
improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
 
wenzelm 
parents: 
61343 
diff
changeset
 | 
16  | 
have 1: "\<exists>x :: 'a. x = x" by (intro exI refl)  | 
| 20005 | 17  | 
|
18  | 
from 1 guess ..  | 
|
19  | 
from 1 guess x ..  | 
|
20  | 
from 1 guess x :: 'a ..  | 
|
| 
74365
 
b49bd5d9041f
improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
 
wenzelm 
parents: 
61343 
diff
changeset
 | 
21  | 
from 1 guess x :: t ..  | 
| 20005 | 22  | 
|
| 
74365
 
b49bd5d9041f
improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
 
wenzelm 
parents: 
61343 
diff
changeset
 | 
23  | 
have 2: "\<exists>(x::'c) (y::'d). x = x \<and> y = y" by (intro exI conjI refl)  | 
| 20005 | 24  | 
from 2 guess apply - apply (erule exE conjE)+ done  | 
25  | 
from 2 guess x apply - apply (erule exE conjE)+ done  | 
|
26  | 
from 2 guess x y apply - apply (erule exE conjE)+ done  | 
|
27  | 
from 2 guess x :: 'a and y :: 'b apply - apply (erule exE conjE)+ done  | 
|
| 
74365
 
b49bd5d9041f
improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
 
wenzelm 
parents: 
61343 
diff
changeset
 | 
28  | 
from 2 guess x y :: t apply - apply (erule exE conjE)+ done  | 
| 59031 | 29  | 
end  | 
| 20005 | 30  | 
|
31  | 
end  |